1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "static_2/static/frees.ma".
17 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
19 (* Main inversion lemmas ****************************************************)
21 theorem frees_mono: ∀f1,L,T. L ⊢ 𝐅+❨T❩ ≘ f1 → ∀f2. L ⊢ 𝐅+❨T❩ ≘ f2 → f1 ≡ f2.
22 #f1 #L #T #H elim H -f1 -L -T
23 [ /3 width=3 by frees_inv_sort, pr_isi_inv_eq_repl/
25 elim (frees_inv_atom … H) -H #f2 #Hf2 #H destruct
26 /4 width=5 by pr_isi_inv_eq_repl, pr_pushs_eq_repl, pr_eq_next/
27 | #f1 #I #L #V #_ #IH #g2 #H elim (frees_inv_pair … H) -H
28 #f2 #Hf2 #H destruct /3 width=5 by pr_eq_next/
29 | #f1 #I #L #Hf1 #g2 #H elim (frees_inv_unit … H) -H
30 #f2 #Hf2 #H destruct /3 width=5 by pr_isi_inv_eq_repl, pr_eq_next/
31 | #f1 #I #L #i #_ #IH #g2 #H elim (frees_inv_lref … H) -H
32 #f2 #Hf2 #H destruct /3 width=5 by pr_eq_push/
33 | /3 width=3 by frees_inv_gref, pr_isi_inv_eq_repl/
34 | #f1V #f1T #f1 #p #I #L #V #T #_ #_ #Hf1 #IHV #IHT #f2 #H elim (frees_inv_bind … H) -H
35 #f2V #f2T #HV #HT #Hf2 @(pr_sor_mono … Hf1) -Hf1
36 /5 width=3 by pr_sor_eq_repl_fwd_dx, pr_sor_eq_repl_fwd_sn, pr_tl_eq_repl/ (**) (* full auto too slow *)
37 | #f1V #f1T #f1 #I #L #V #T #_ #_ #Hf1 #IHV #IHT #f2 #H elim (frees_inv_flat … H) -H
38 #f2V #f2T #HV #HT #Hf2 @(pr_sor_mono … Hf1) -Hf1
39 /4 width=3 by pr_sor_eq_repl_fwd_dx, pr_sor_eq_repl_fwd_sn/ (**) (* full auto too slow *)