]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/relocation/lleq_lleq.ma
- lambdadelta: first commutation property on lazy equivalence for
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / lleq_lleq.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "basic_2/relocation/ldrop_ldrop.ma".
16 include "basic_2/relocation/lleq.ma".
17
18 (* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************)
19
20 (* Advanced inversion lemmas ************************************************)
21
22 lemma lleq_inv_lref_dx: ∀L1,L2,i. L1 ⋕[#i] L2 →
23                         ∀I,K2,V. ⇩[0, i] L2 ≡ K2.ⓑ{I}V →
24                         ∃∃K1. ⇩[0, i] L1 ≡ K1.ⓑ{I}V & K1 ⋕[V] K2.
25 #L1 #L2 #i #H #I #K2 #V #HLK2 elim (lleq_inv_lref … H) -H *
26 [ #_ #H elim (lt_refl_false i)
27   /3 width=5 by ldrop_fwd_length_lt2, lt_to_le_to_lt/
28 | #I0 #K1 #K0 #V0 #HLK1 #HLK0 #HK10 lapply (ldrop_mono … HLK0 … HLK2) -L2
29   #H destruct /2 width=3 by ex2_intro/
30 ]
31 qed-.
32
33 lemma lleq_inv_lift: ∀L1,L2,U. L1 ⋕[U] L2 →
34                      ∀K1,K2,d,e. ⇩[d, e] L1 ≡ K1 → ⇩[d, e] L2 ≡ K2 →
35                      ∀T. ⇧[d, e] T ≡ U → K1 ⋕[T] K2.
36 #L1 #L2 #U #H elim H -L1 -L2 -U
37 [ #L1 #L2 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H >(lift_inv_sort2 … H) -X
38   lapply (ldrop_fwd_length_eq … HLK1 HLK2 HL12) -L1 -L2 -d -e
39   /2 width=1 by lleq_sort/ (**) (* full auto fails *)
40 | #I #L1 #L2 #K #K0 #W #i #HLK #HLK0 #HK0 #IHK0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_lref2 … H) -H
41   * #Hdei #H destruct [ -HK0 | -IHK0 ]
42   [ elim (ldrop_conf_lt … HLK1 … HLK) // -L1 #L1 #V #HKL1 #KL1 #HV0
43     elim (ldrop_conf_lt … HLK2 … HLK0) // -Hdei -L2 #L2 #V2 #HKL2 #K0L2 #HV02
44     lapply (lift_inj … HV02 … HV0) -HV02 #H destruct /3 width=11 by lleq_lref/
45   | lapply (ldrop_conf_ge … HLK1 … HLK ?) // -L1
46     lapply (ldrop_conf_ge … HLK2 … HLK0 ?) // -Hdei -L2
47     /2 width=7 by lleq_lref/
48   ]
49 | #L1 #L2 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_lref2 … H) -H
50   * #_ #H destruct
51   lapply (ldrop_fwd_length_eq … HLK1 HLK2 HL12)
52   [ /4 width=3 by lleq_free, ldrop_fwd_length_le4, transitive_le/
53   | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H
54     lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H
55     /3 width=1 by lleq_free, le_plus_to_minus_r/
56   ]
57 | #L1 #L2 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H >(lift_inv_gref2 … H) -X
58   lapply (ldrop_fwd_length_eq … HLK1 HLK2 HL12) -L1 -L2 -d -e
59   /2 width=1 by lleq_gref/ (**) (* full auto fails *)
60 | #a #I #L1 #L2 #W #U #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_bind2 … H) -H
61   #V #T #HVW #HTU #H destruct /4 width=5 by lleq_bind, ldrop_skip/
62 | #I #L1 #L2 #W #U #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H elim (lift_inv_flat2 … H) -H
63   #V #T #HVW #HTU #H destruct /3 width=5 by lleq_flat/
64 ]
65 qed-.
66
67 (* Advanced properties ******************************************************)
68
69 lemma lleq_dec: ∀T,L1,L2. Decidable (L1 ⋕[T] L2).
70 #T #L1 @(f2_ind … rfw … L1 T) -L1 -T
71 #n #IH #L1 * *
72 [ #k #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_sort/
73 | #i #H1 #L2 elim (eq_nat_dec (|L1|) (|L2|))
74   [ #HL12 elim (lt_or_ge i (|L1|))
75     #HiL1 elim (lt_or_ge i (|L2|)) /3 width=1 by or_introl, lleq_free/
76     #HiL2 elim (ldrop_O1_lt … HiL2)
77     #I2 #K2 #V2 #HLK2 elim (ldrop_O1_lt … HiL1)
78     #I1 #K1 #V1 #HLK1 elim (eq_bind2_dec I2 I1)
79     [ #H2 elim (eq_term_dec V2 V1)
80       [ #H3 elim (IH K1 V1 … K2) destruct
81         /3 width=7 by lleq_lref, ldrop_fwd_rfw, or_introl/
82       ]
83     ]
84     -IH #H3 @or_intror
85     #H elim (lleq_inv_lref … H) -H *
86     [1,3,5: /3 width=4 by lt_to_le_to_lt, lt_refl_false/ ]
87     #I0 #X1 #X2 #V0 #HLX1 #HLX2 #HX12
88     lapply (ldrop_mono … HLX1 … HLK1) -HLX1 -HLK1
89     lapply (ldrop_mono … HLX2 … HLK2) -HLX2 -HLK2
90     #H #H0 destruct /2 width=1 by/
91   ]
92 | #p #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_gref/
93 | #a #I #V #T #Hn #L2 destruct
94   elim (IH L1 V … L2) /2 width=1 by/
95   elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V)) -IH /3 width=1 by or_introl, lleq_bind/
96   #H1 #H2 @or_intror
97   #H elim (lleq_inv_bind … H) -H /2 width=1 by/
98 | #I #V #T #Hn #L2 destruct
99   elim (IH L1 V … L2) /2 width=1 by/
100   elim (IH L1 T … L2) -IH /3 width=1 by or_introl, lleq_flat/
101   #H1 #H2 @or_intror
102   #H elim (lleq_inv_flat … H) -H /2 width=1 by/
103 ]
104 -n /4 width=2 by lleq_fwd_length, or_intror/
105 qed-.
106
107 (* Main properties **********************************************************)
108
109 theorem lleq_trans: ∀T. Transitive … (lleq T).
110 #T #L1 #L #H elim H -T -L1 -L
111 /4 width=4 by lleq_fwd_length, lleq_gref, lleq_sort, trans_eq/
112 [ #I #L1 #L #K1 #K #V #i #HLK1 #HLK #_ #IHK1 #L2 #H elim (lleq_inv_lref … H) -H *
113   [ -HLK1 -IHK1 #HLi #_ elim (lt_refl_false i)
114     /3 width=5 by ldrop_fwd_length_lt2, lt_to_le_to_lt/
115   | #I0 #K0 #K2 #V0 #HLK0 #HLK2 #HK12 lapply (ldrop_mono … HLK0 … HLK) -L
116     #H destruct /3 width=7 by lleq_lref/
117   ]
118 | #L1 #L #i #HL1i #HLi #HL #L2 #H lapply (lleq_fwd_length … H)
119   #HL2 elim (lleq_inv_lref … H) -H * /2 width=1 by lleq_free/
120 | #a #I #L1 #L #V #T #_ #_ #IHV #IHT #L2 #H elim (lleq_inv_bind … H) -H
121   /3 width=1 by lleq_bind/
122 | #I #L1 #L #V #T #_ #_ #IHV #IHT #L2 #H elim (lleq_inv_flat … H) -H
123   /3 width=1 by lleq_flat/
124 ]
125 qed-.
126
127 theorem lleq_canc_sn: ∀L,L1,L2,T. L ⋕[T] L1→ L ⋕[T] L2 → L1 ⋕[T] L2.
128 /3 width=3 by lleq_trans, lleq_sym/ qed-.
129
130 theorem lleq_canc_dx: ∀L1,L2,L,T. L1 ⋕[T] L → L2 ⋕[T] L → L1 ⋕[T] L2.
131 /3 width=3 by lleq_trans, lleq_sym/ qed-.