]> matita.cs.unibo.it Git - helm.git/commitdiff
update in basic_2
authorFerruccio Guidi <fguidi@maelstrom.helm.cs.unibo.it>
Wed, 17 Jan 2018 19:31:56 +0000 (20:31 +0100)
committerFerruccio Guidi <fguidi@maelstrom.helm.cs.unibo.it>
Wed, 17 Jan 2018 19:31:56 +0000 (20:31 +0100)
yet another definition of lveq allows to prove a missing lemma in lfxs_fle

.gitignore
matita/matita/contribs/lambdadelta/basic_2/static/fle.ma
matita/matita/contribs/lambdadelta/basic_2/static/fle_drops.ma
matita/matita/contribs/lambdadelta/basic_2/static/fle_fle.ma
matita/matita/contribs/lambdadelta/basic_2/static/fle_fqup.ma
matita/matita/contribs/lambdadelta/basic_2/static/lfxs_fle.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/lenv_length.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/lveq.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/lveq_length.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/lveq_lveq.ma
matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl

index 7ab34bb92666be35cb819c59392ed1dbd537008c..45d3126b7e85a9a80bf49b6681cadb910255439b 100644 (file)
@@ -66,6 +66,7 @@ helm/www/lambdadelta/xslt/core.xsl
 helm/www/lambdadelta/xslt/chc_45.xsl
 helm/www/lambdadelta/xslt/xhtbl.xsl
 
+matita/matita/contribs/lambdadelta/nodes
 matita/matita/contribs/lambdadelta/token
 matita/matita/contribs/lambdadelta/2A1
 matita/matita/contribs/lambdadelta/apps_2/notation
index fdc03a5e2388869cc71ec99515f35d0ece4c1017..95d12b132fd0efe0ad5728b65c3242689890b788 100644 (file)
@@ -29,14 +29,10 @@ interpretation "free variables inclusion (restricted closure)"
 (* Basic properties *********************************************************)
 
 lemma fle_sort: ∀L,s1,s2. ⦃L, ⋆s1⦄ ⊆ ⦃L, ⋆s2⦄.
-#L elim (lveq_refl L)
-/3 width=8 by frees_sort, sle_refl, ex4_4_intro/
-qed.
+/3 width=8 by frees_sort, sle_refl, ex4_4_intro/ qed.
 
 lemma fle_gref: ∀L,l1,l2. ⦃L, §l1⦄ ⊆ ⦃L, §l2⦄.
-#L elim (lveq_refl L)
-/3 width=8 by frees_gref, sle_refl, ex4_4_intro/
-qed.
+/3 width=8 by frees_gref, sle_refl, ex4_4_intro/ qed.
 
 (* Basic inversion lemmas ***************************************************)
 (*
index ff0b8872d6f403b0a2e7e8c6295e3734c57aee63..4ce02b60b6e19d9304e0797d52502770f7828589 100644 (file)
@@ -26,6 +26,5 @@ elim (frees_total L V) #f1 #Hf1
 elim (frees_total L T) #f2 #Hf2
 elim (sor_isfin_ex f1 f2) /3 width=3 by frees_fwd_isfin, isfin_tl/ #f #Hf #_
 lapply (sor_inv_sle_dx … Hf) #Hf0
-elim (lveq_refl L) #n #HL
 /6 width=8 by frees_lifts_SO, frees_bind, drops_refl, drops_drop, sle_tls, ex4_4_intro/
 qed.
index 42c403412849db378d75a32747c5a0a34bd6052e..4301456f5871f052dd4409d823b138ea6ab17218 100644 (file)
@@ -38,11 +38,12 @@ theorem fle_trans: bi_transitive … fle.
 /5 width=8 by frees_mono, sle_trans, sle_eq_repl_back2, ex3_2_intro/
 qed-.
 *)
-theorem fle_bind_sn: ∀L1,L2,V1,T1,T. ⦃L1, V1⦄ ⊆ ⦃L2, T⦄ → ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2, T⦄ →
-                     ∀p,I. ⦃L1, ⓑ{p,I}V1.T1⦄ ⊆ ⦃L2, T⦄.
-#L1 #L2 #V1 #T1 #T * #n1 #x #f1 #g #Hf1 #Hg #H1n1 #H2n1 #H #p #I
+theorem fle_bind_sn_ge: ∀L1,L2. |L2| ≤ |L1| →
+                        ∀V1,T1,T. ⦃L1, V1⦄ ⊆ ⦃L2, T⦄ → ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2, T⦄ →
+                        ∀p,I. ⦃L1, ⓑ{p,I}V1.T1⦄ ⊆ ⦃L2, T⦄.
+#L1 #L2 #HL #V1 #T1 #T * #n1 #x #f1 #g #Hf1 #Hg #H1n1 #H2n1 #H #p #I
 elim (fle_frees_trans … H … Hg) -H #n2 #n #f2 #Hf2 #H1n2 #H2n2
-elim (lveq_inj_void_sn … H1n1 … H1n2) -H1n2 #H1 #H2 destruct
+elim (lveq_inj_void_sn_ge … H1n1 … H1n2) -H1n2 // #H1 #H2 #H3 destruct
 elim (sor_isfin_ex f1 (⫱f2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #f #Hf #_
 <tls_xn in H2n2; #H2n2
 /4 width=12 by frees_bind_void, sor_inv_sle, sor_tls, ex4_4_intro/
index 4db4e617c829dcb7bd408e653e294c5391915088..65237fbc1d60d4d617f5155ad3be6b0d295cc54f 100644 (file)
@@ -22,7 +22,6 @@ include "basic_2/static/fle.ma".
 
 lemma fle_refl: bi_reflexive … fle.
 #L #T
-elim (lveq_refl L) #n #Hn
 elim (frees_total L T) #f #Hf
 /2 width=8 by sle_refl, ex4_4_intro/
 qed.
@@ -39,7 +38,7 @@ qed.
 lemma fle_bind_dx_dx: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L2.ⓧ, T2⦄ → |L1| ≤ |L2| →
                       ∀p,I,V2. ⦃L1, T1⦄ ⊆ ⦃L2, ⓑ{p,I}V2.T2⦄.
 #L1 #L2 #T1 #T2 * #n1 #x1 #f2 #g2 #Hf2 #Hg2 #H #Hfg2 #HL12 #p #I #V2
-elim (lveq_inv_void_dx_length … H HL12) -H -HL12 #m1 #HL12 #H #_ destruct
+elim (lveq_inv_void_dx_length … H HL12) -H -HL12 #m1 #HL12 #H1 #H2 destruct
 <tls_xn in Hfg2; #Hfg2
 elim (frees_total L2 V2) #g1 #Hg1
 elim (sor_isfin_ex g1 (⫱g2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #g #Hg #_
index c13c4dcec51a08fc69547ddd29755f30cc2e7286..cc067e50ae7ac5668174a7149b79b8189b688a03 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-(* include "basic_2/syntax/lveq_length.ma". *)
+include "basic_2/syntax/lveq_length.ma".
 include "basic_2/static/fle.ma".
 include "basic_2/static/lfxs_lfxs.ma".
 
 (* GENERIC EXTENSION ON REFERRED ENTRIES OF A CONTEXT-SENSITIVE REALTION ****)
-(*
-lemma pippo: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-             ∀T,f. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∃g. ↑*[n1]g = f. 
-*)
+
 (* Properties with free variables inclusion for restricted closures *********)
-(*
+
+(* Note: we just need lveq_inv_refl: ∀L,n1,n2. L ≋ⓧ*[n1, n2] L → ∧∧ 0 = n1 & 0 = n2 *)
 lemma fle_lfxs_trans: ∀R,L1,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L1, T2⦄ →
                       ∀L2. L1 ⪤*[R, T2] L2 → L1 ⪤*[R, T1] L2.
-#R #L1 #T1 #T2 * #x #n #f1 #f2 #Hf1 #Hf2 #Hn #Hf #L2 #HL12
-lapply (lveq_inj_length … Hn ?) // #H destruct
-
-
- Hn : (L1≋ⓧ*[n,n]L1) (L1⊢𝐅*⦃T1⦄≡f1) → 
-
-lapply (lfxs_inv_frees … HL12 … Hf2) -HL12 -Hf2 #HL12
-@(ex2_intro … Hf1) -Hf1
-@(sle_lexs_trans … HL12) -HL12 //
-
+#R #L1 #T1 #T2 * #n1 #n2 #f1 #f2 #Hf1 #Hf2 #Hn #Hf #L2 #HL12
+elim (lveq_inj_length … Hn ?) // #H1 #H2 destruct
 /4 width=5 by lfxs_inv_frees, sle_lexs_trans, ex2_intro/
 qed-.
-*)
\ No newline at end of file
index 8b7ccd58207a7b1fc92fc4203da6f70612a7982c..d2690db960f24ca0950f85668e56106ff582e1f7 100644 (file)
@@ -23,8 +23,6 @@ rec definition length L ≝ match L with
 
 interpretation "length (local environment)" 'card L = (length L).
 
-definition length2 (L1) (L2): nat ≝ |L1| + |L2|. 
-
 (* Basic properties *********************************************************)
 
 lemma length_atom: |⋆| = 0.
index c9300df18d65ed40df140f5c58cb6a068f512b54..e19708385f8b7d4fbcfdff2d99a4e7a5844adae3 100644 (file)
@@ -19,14 +19,12 @@ include "basic_2/syntax/lenv.ma".
 
 inductive lveq: bi_relation nat lenv ≝
 | lveq_atom   : lveq 0 (⋆) 0 (⋆)
-| lveq_pair_sn: ∀I1,I2,K1,K2,V1,n. lveq n K1 n K2 →
-                lveq 0 (K1.ⓑ{I1}V1) 0 (K2.ⓘ{I2})
-| lveq_pair_dx: ∀I1,I2,K1,K2,V2,n. lveq n K1 n K2 →
-                lveq 0 (K1.ⓘ{I1}) 0 (K2.ⓑ{I2}V2)
-| lveq_void_sn: ∀K1,K2,n1,n2. lveq n1 K1 n2 K2 →
-                lveq (⫯n1) (K1.ⓧ) n2 K2
-| lveq_void_dx: ∀K1,K2,n1,n2. lveq n1 K1 n2 K2 →
-                lveq n1 K1 (⫯n2) (K2.ⓧ)
+| lveq_bind   : ∀I1,I2,K1,K2. lveq 0 K1 0 K2 →
+                lveq 0 (K1.ⓘ{I1}) 0 (K2.ⓘ{I2})
+| lveq_void_sn: ∀K1,K2,n1. lveq n1 K1 0 K2 →
+                lveq (⫯n1) (K1.ⓧ) 0 K2
+| lveq_void_dx: ∀K1,K2,n2. lveq 0 K1 n2 K2 →
+                lveq 0 K1 (⫯n2) (K2.ⓧ)
 .
 
 interpretation "equivalence up to exclusion binders (local environment)"
@@ -34,113 +32,140 @@ interpretation "equivalence up to exclusion binders (local environment)"
 
 (* Basic properties *********************************************************)
 
-lemma lveq_refl: ∀L. ∃n. L ≋ⓧ*[n, n] L.
-#L elim L -L /2 width=2 by ex_intro, lveq_atom/
-#L #I * #n #IH cases I -I /3 width=2 by ex_intro, lveq_pair_dx/
-* /4 width=2 by ex_intro, lveq_void_sn, lveq_void_dx/
-qed-.
+lemma lveq_refl: ∀L. L ≋ⓧ*[0, 0] L.
+#L elim L -L /2 width=1 by lveq_atom, lveq_bind/
+qed.
 
 lemma lveq_sym: bi_symmetric … lveq.
 #n1 #n2 #L1 #L2 #H elim H -L1 -L2 -n1 -n2
-/2 width=2 by lveq_atom, lveq_pair_sn, lveq_pair_dx, lveq_void_sn, lveq_void_dx/
+/2 width=1 by lveq_atom, lveq_bind, lveq_void_sn, lveq_void_dx/
 qed-.
 
 (* Basic inversion lemmas ***************************************************)
 
-fact lveq_inv_atom_atom_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-                             ⋆ = L1 → ⋆ = L2 → ∧∧ 0 = n1 & 0 = n2.
+fact lveq_inv_zero_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
+                        0 = n1 → 0 = n2 →
+                        ∨∨ ∧∧ ⋆ = L1 & ⋆ = L2
+                            | ∃∃I1,I2,K1,K2. K1 ≋ⓧ*[0, 0] K2 & K1.ⓘ{I1} = L1 & K2.ⓘ{I2} = L2.
 #L1 #L2 #n1 #n2 * -L1 -L2 -n1 -n2
-[ /2 width=1 by conj/
-|2,3: #I1 #I2 #K1 #K2 #V #n #_ #H1 #H2 destruct
-|4,5: #K1 #K2 #n1 #n2 #_ #H1 #H2 destruct
+[1: /3 width=1 by or_introl, conj/
+|2: /3 width=7 by ex3_4_intro, or_intror/
+|*: #K1 #K2 #n #_ #H1 #H2 destruct
 ]
 qed-.
 
-lemma lveq_inv_atom_atom: ∀n1,n2. ⋆ ≋ⓧ*[n1, n2] ⋆ → 0 = n1 ∧ 0 = n2.
-/2 width=5 by lveq_inv_atom_atom_aux/ qed-.
+lemma lveq_inv_zero: ∀L1,L2. L1 ≋ⓧ*[0, 0] L2 →
+                     ∨∨ ∧∧ ⋆ = L1 & ⋆ = L2
+                      | ∃∃I1,I2,K1,K2. K1 ≋ⓧ*[0, 0] K2 & K1.ⓘ{I1} = L1 & K2.ⓘ{I2} = L2.
+/2 width=5 by lveq_inv_zero_aux/ qed-.
 
-fact lveq_inv_bind_atom_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-                             ∀I1,K1. K1.ⓘ{I1} = L1 → ⋆ = L2 →
-                             ∃∃m1. K1 ≋ⓧ*[m1, n2] ⋆ & BUnit Void = I1 & ⫯m1 = n1.
+fact lveq_inv_succ_sn_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
+                           ∀m1. ⫯m1 = n1 →
+                           ∃∃K1. K1 ≋ⓧ*[m1, 0] L2 & K1.ⓧ = L1 & 0 = n2.
 #L1 #L2 #n1 #n2 * -L1 -L2 -n1 -n2
-[ #Z1 #Y1 #H destruct
-|2,3: #I1 #I2 #K1 #K2 #V #n #_ #Z1 #Y1 #_ #H2 destruct
-|4,5: #K1 #K2 #n1 #n2 #HK #Z1 #Y1 #H1 #H2 destruct /2 width=3 by ex3_intro/
+[1: #m #H destruct
+|2: #I1 #I2 #K1 #K2 #_ #m #H destruct
+|*: #K1 #K2 #n #HK #m #H destruct /2 width=3 by ex3_intro/
 ]
 qed-.
 
-lemma lveq_inv_bind_atom: ∀I1,K1,n1,n2. K1.ⓘ{I1} ≋ⓧ*[n1, n2] ⋆ →
-                          ∃∃m1. K1 ≋ⓧ*[m1, n2] ⋆ & BUnit Void = I1 & ⫯m1 = n1.
-/2 width=5 by lveq_inv_bind_atom_aux/ qed-.
+lemma lveq_inv_succ_sn: ∀L1,K2,n1,n2. L1 ≋ⓧ*[⫯n1, n2] K2 →
+                        ∃∃K1. K1 ≋ⓧ*[n1, 0] K2 & K1.ⓧ = L1 & 0 = n2.
+/2 width=3 by lveq_inv_succ_sn_aux/ qed-.
 
-lemma lveq_inv_atom_bind: ∀I2,K2,n1,n2. ⋆ ≋ⓧ*[n1, n2] K2.ⓘ{I2} →
-                          ∃∃m2. ⋆ ≋ⓧ*[n1, m2] K2 & BUnit Void = I2 & ⫯m2 = n2.
-#I2 #K2 #n1 #n2 #H
+lemma lveq_inv_succ_dx: ∀K1,L2,n1,n2. K1 ≋ⓧ*[n1, ⫯n2] L2 →
+                        ∃∃K2. K1 ≋ⓧ*[0, n2] K2 & K2.ⓧ = L2 & 0 = n1.
+#K1 #L2 #n1 #n2 #H
 lapply (lveq_sym … H) -H #H
-elim (lveq_inv_bind_atom … H) -H
-/3 width=3 by lveq_sym, ex3_intro/
+elim (lveq_inv_succ_sn … H) -H /3 width=3 by lveq_sym, ex3_intro/
 qed-.
 
-fact lveq_inv_pair_pair_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-                             ∀I1,I2,K1,K2,V1,V2. K1.ⓑ{I1}V1 = L1 → K2.ⓑ{I2}V2 = L2 →
-                             ∃∃n. K1 ≋ⓧ*[n, n] K2 & 0 = n1 & 0 = n2.
+fact lveq_inv_succ_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
+                        ∀m1,m2. ⫯m1 = n1 → ⫯m2 = n2 → ⊥.
 #L1 #L2 #n1 #n2 * -L1 -L2 -n1 -n2
-[ #Z1 #Z2 #Y1 #Y2 #X1 #X2 #H1 #H2 destruct
-|2,3: #I1 #I2 #K1 #K2 #V #n #HK #Z1 #Z2 #Y1 #Y2 #X1 #X2 #H1 #H2 destruct /2 width=2 by ex3_intro/
-|4,5: #K1 #K2 #n1 #n2 #_ #Z1 #Z2 #Y1 #Y2 #X1 #X2 #H1 #H2 destruct
+[1: #m1 #m2 #H1 #H2 destruct
+|2: #I1 #I2 #K1 #K2 #_ #m1 #m2 #H1 #H2 destruct
+|*: #K1 #K2 #n #_ #m1 #m2 #H1 #H2 destruct
 ]
 qed-.
 
-lemma lveq_inv_pair_pair: ∀I1,I2,K1,K2,V1,V2,m1,m2. K1.ⓑ{I1}V1 ≋ⓧ*[m1, m2] K2.ⓑ{I2}V2 →
-                          ∃∃n. K1 ≋ⓧ*[n, n] K2 & 0 = m1 & 0 = m2.
-/2 width=9 by lveq_inv_pair_pair_aux/ qed-.
+lemma lveq_inv_succ: ∀L1,L2,n1,n2. L1 ≋ⓧ*[⫯n1, ⫯n2] L2 → ⊥.
+/2 width=9 by lveq_inv_succ_aux/ qed-.
 
 (* Advanced inversion lemmas ************************************************)
 
-fact lveq_inv_void_succ_sn_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-                                ∀K1,m1. L1 = K1.ⓧ → n1 = ⫯m1 → K1 ≋ ⓧ*[m1, n2] L2.
-#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2
-[ #K2 #m2 #H destruct
-| #I1 #I2 #L1 #L2 #V #n #_ #_ #K1 #m1 #H1 #H2 destruct
-| #I1 #I2 #L1 #L2 #V #n #_ #_ #K1 #m1 #H1 #H2 destruct
-| #L1 #L2 #n1 #n2 #HL12 #_ #K1 #m1 #H1 #H2 destruct //
-| #L1 #L2 #n1 #n2 #_ #IH #K1 #m1 #H1 #H2 destruct
-  /3 width=1 by lveq_void_dx/
+lemma lveq_inv_atom_atom: ∀n1,n2. ⋆ ≋ⓧ*[n1, n2] ⋆ → ∧∧ 0 = n1 & 0 = n2.
+* [2: #n1 ] * [2,4: #n2 ] #H
+[ elim (lveq_inv_succ … H)
+| elim (lveq_inv_succ_dx … H) -H #Y #_ #H1 #H2 destruct
+| elim (lveq_inv_succ_sn … H) -H #Y #_ #H1 #H2 destruct
+| /2 width=1 by conj/
 ]
 qed-.
 
-lemma lveq_inv_void_succ_sn: ∀L1,L2,n1,n2. L1.ⓧ ≋ⓧ*[⫯n1, n2] L2 → L1 ≋ ⓧ*[n1, n2] L2.
-/2 width=5 by lveq_inv_void_succ_sn_aux/ qed-.
-
-lemma lveq_inv_void_succ_dx: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, ⫯n2] L2.ⓧ → L1 ≋ ⓧ*[n1, n2] L2.
-/4 width=5 by lveq_inv_void_succ_sn_aux, lveq_sym/ qed-.
+lemma lveq_inv_bind_atom: ∀I1,K1,n1,n2. K1.ⓘ{I1} ≋ⓧ*[n1, n2] ⋆ →
+                          ∃∃m1. K1 ≋ⓧ*[m1, 0] ⋆ & BUnit Void = I1 & ⫯m1 = n1 & 0 = n2.
+#I1 #K1 * [2: #n1 ] * [2,4: #n2 ] #H
+[ elim (lveq_inv_succ … H)
+| elim (lveq_inv_succ_dx … H) -H #Y #_ #H1 #H2 destruct
+| elim (lveq_inv_succ_sn … H) -H #Y #HY #H1 #H2 destruct /2 width=3 by ex4_intro/
+| elim (lveq_inv_zero … H) -H *
+  [ #H1 #H2 destruct
+  | #Z1 #Z2 #Y1 #Y2 #_ #H1 #H2 destruct
+  ]
+]
+qed-.
 
-(* Basic forward lemmas *****************************************************)
+lemma lveq_inv_atom_bind: ∀I2,K2,n1,n2. ⋆ ≋ⓧ*[n1, n2] K2.ⓘ{I2} →
+                          ∃∃m2. ⋆ ≋ⓧ*[0, m2] K2 & BUnit Void = I2 & 0 = n1 & ⫯m2 = n2.
+#I2 #K2 #n1 #n2 #H
+lapply (lveq_sym … H) -H #H
+elim (lveq_inv_bind_atom … H) -H
+/3 width=3 by lveq_sym, ex4_intro/
+qed-.
 
-fact lveq_fwd_void_void_aux: ∀L1,L2,m1,m2. L1 ≋ⓧ*[m1, m2] L2 →
-                             ∀K1,K2. K1.ⓧ = L1 → K2.ⓧ = L2 →
-                             ∨∨ ∃n. ⫯n = m1 | ∃n. ⫯n = m2.
-#L1 #L2 #m1 #m2 * -L1 -L2 -m1 -m2
-[ #Y1 #Y2 #H1 #H2 destruct
-|2,3: #I1 #I2 #K1 #K2 #V #n #_ #Y1 #Y2 #H1 #H2 destruct
-|4,5: #K1 #K2 #n1 #n2 #_ #Y1 #Y2 #H1 #H2 destruct /3 width=2 by ex_intro, or_introl, or_intror/
+lemma lveq_inv_pair_pair: ∀I1,I2,K1,K2,V1,V2,n1,n2. K1.ⓑ{I1}V1 ≋ⓧ*[n1, n2] K2.ⓑ{I2}V2 →
+                          ∧∧ K1 ≋ⓧ*[0, 0] K2 & 0 = n1 & 0 = n2.
+#I1 #I2 #K1 #K2 #V1 #V2 * [2: #n1 ] * [2,4: #n2 ] #H
+[ elim (lveq_inv_succ … H)
+| elim (lveq_inv_succ_dx … H) -H #Y #_ #H1 #H2 destruct
+| elim (lveq_inv_succ_sn … H) -H #Y #_ #H1 #H2 destruct
+| elim (lveq_inv_zero … H) -H *
+  [ #H1 #H2 destruct
+  | #Z1 #Z2 #Y1 #Y2 #HY #H1 #H2 destruct /3 width=1 by and3_intro/
+  ]
 ]
 qed-.
 
-lemma lveq_fwd_void_void: ∀K1,K2,m1,m2. K1.ⓧ ≋ⓧ*[m1, m2] K2.ⓧ →
-                          ∨∨ ∃n. ⫯n = m1 | ∃n. ⫯n = m2.
-/2 width=7 by lveq_fwd_void_void_aux/ qed-.
+lemma lveq_inv_void_succ_sn: ∀L1,L2,n1,n2. L1.ⓧ ≋ⓧ*[⫯n1, n2] L2 →
+                             ∧∧ L1 ≋ ⓧ*[n1, 0] L2 & 0 = n2.
+#L1 #L2 #n1 #n2 #H
+elim (lveq_inv_succ_sn … H) -H #Y #HY #H1 #H2 destruct /2 width=1 by conj/
+qed-.
+
+lemma lveq_inv_void_succ_dx: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, ⫯n2] L2.ⓧ →
+                             ∧∧ L1 ≋ ⓧ*[0, n2] L2 & 0 = n1.
+#L1 #L2 #n1 #n2 #H
+lapply (lveq_sym … H) -H #H
+elim (lveq_inv_void_succ_sn … H) -H
+/3 width=1 by lveq_sym, conj/
+qed-.
 
 (* Advanced forward lemmas **************************************************)
 
-fact lveq_fwd_pair_sn_aux: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-                           ∀I,K1,V. K1.ⓑ{I}V = L1 → 0 = n1.
-#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 //
-#K1 #K2 #n1 #n2 #_ #IH #J #L1 #V #H destruct /2 width=4 by/
+lemma lveq_fwd_gen: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
+                    ∨∨ 0 = n1 | 0 = n2.
+#L1 #L2 * [2: #n1 ] * [2,4: #n2 ] #H
+[ elim (lveq_inv_succ … H) ]
+/2 width=1 by or_introl, or_intror/
 qed-.
 
-lemma lveq_fwd_pair_sn: ∀I,K1,L2,V,n1,n2. K1.ⓑ{I}V ≋ⓧ*[n1, n2] L2 → 0 = n1.
-/2 width=8 by lveq_fwd_pair_sn_aux/ qed-.
+lemma lveq_fwd_pair_sn: ∀I1,K1,L2,V1,n1,n2. K1.ⓑ{I1}V1 ≋ⓧ*[n1, n2] L2 → 0 = n1.
+#I1 #K1 #L2 #V1 * [2: #n1 ] // * [2: #n2 ] #H
+[ elim (lveq_inv_succ … H)
+| elim (lveq_inv_succ_sn … H) -H #Y #_ #H1 #H2 destruct
+]
+qed-.
 
-lemma lveq_fwd_pair_dx: ∀I,L1,K2,V,n1,n2. L1 ≋ⓧ*[n1, n2] K2.ⓑ{I}V → 0 = n2.
+lemma lveq_fwd_pair_dx: ∀I2,L1,K2,V2,n1,n2. L1 ≋ⓧ*[n1, n2] K2.ⓑ{I2}V2 → 0 = n2.
 /3 width=6 by lveq_fwd_pair_sn, lveq_sym/ qed-.
index 3695d2a7cc9fa073213cb55d448da5c9990c3e4f..ab344552272c1cd9c57b8989a5793e2c7e9e7b7e 100644 (file)
@@ -17,15 +17,16 @@ include "basic_2/syntax/lveq.ma".
 
 (* EQUIVALENCE FOR LOCAL ENVIRONMENTS UP TO EXCLUSION BINDERS ***************)
 
-lemma lveq_eq_ex: ∀L1,L2. |L1| = |L2| → ∃n. L1 ≋ⓧ*[n, n] L2.
+(* Properties with length for local environments ****************************)
+
+lemma lveq_length_eq: ∀L1,L2. |L1| = |L2| → L1 ≋ⓧ*[0, 0] L2.
 #L1 elim L1 -L1
 [ #Y2 #H >(length_inv_zero_sn … H) -Y2 /2 width=3 by lveq_atom, ex_intro/
-| #K1 * [ * | #I1 #V1 ] #IH #Y2 #H
-  elim (length_inv_succ_sn … H) -H * [1,3: * |*: #I2 #V2 ] #K2 #HK #H destruct 
-  elim (IH … HK) -IH -HK #n #HK
-  /4 width=3 by lveq_pair_sn, lveq_pair_dx, lveq_void_sn, lveq_void_dx, ex_intro/
+| #K1 #I1 #IH #Y2 #H
+  elim (length_inv_succ_sn … H) -H #I2 #K2 #HK #H destruct
+  /3 width=1 by lveq_bind/
 ]
-qed-.
+qed.
 
 (* Forward lemmas with length for local environments ************************)
 
@@ -40,31 +41,36 @@ lemma lveq_fwd_length_le_dx: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → n2 ≤ |L
 qed-.
 
 lemma lveq_fwd_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-                       |L1| + n2 = |L2| + n1.
+                       ∧∧ |L1|-|L2| = n1 & |L2|-|L1| = n2.
+#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 /2 width=1 by conj/
+#K1 #K2 #n #_ * #H1 #H2 >length_bind /3 width=1 by minus_Sn_m, conj/
+qed-.
+
+lemma lveq_inj_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
+                       |L1| = |L2| → ∧∧ 0 = n1 & 0 = n2.
+#L1 #L2 #n1 #n2 #H #HL
+elim (lveq_fwd_length … H) -H
+>HL -HL /2 width=1 by conj/ 
+qed-.
+
+lemma lveq_fwd_length_plus: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
+                            |L1| + n2 = |L2| + n1.
 #L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 normalize
 /2 width=2 by injective_plus_r/
 qed-.
 
-lemma lveq_fwd_length_eq: ∀L1,L2,n. L1 ≋ⓧ*[n, n] L2 → |L1| = |L2|.
-/3 width=2 by lveq_fwd_length, injective_plus_l/ qed-.
+lemma lveq_fwd_length_eq: ∀L1,L2. L1 ≋ⓧ*[0, 0] L2 → |L1| = |L2|.
+/3 width=2 by lveq_fwd_length_plus, injective_plus_l/ qed-.
 
 lemma lveq_fwd_length_minus: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
                              |L1| - n1 = |L2| - n2.
-/3 width=3 by lveq_fwd_length, lveq_fwd_length_le_dx, lveq_fwd_length_le_sn, plus_to_minus_2/ qed-.
-
-lemma lveq_inj_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
-                       |L1| = |L2| → n1 = n2.
-#L1 #L2 #n1 #n2 #H #HL12
-lapply (lveq_fwd_length … H) -H #H
-/2 width=2 by injective_plus_l/
-qed-.
+/3 width=3 by lveq_fwd_length_plus, lveq_fwd_length_le_dx, lveq_fwd_length_le_sn, plus_to_minus_2/ qed-.
 
 lemma lveq_fwd_abst_bind_length_le: ∀I1,I2,L1,L2,V1,n1,n2.
                                     L1.ⓑ{I1}V1 ≋ⓧ*[n1, n2] L2.ⓘ{I2} → |L1| ≤ |L2|.
 #I1 #I2 #L1 #L2 #V1 #n1 #n2 #HL
 lapply (lveq_fwd_pair_sn … HL) #H destruct
-lapply (lveq_fwd_length … HL) -HL >length_bind >length_bind #H
-/2 width=1 by monotonic_pred/
+elim (lveq_fwd_length … HL) -HL >length_bind >length_bind //
 qed-.
 
 lemma lveq_fwd_bind_abst_length_le: ∀I1,I2,L1,L2,V2,n1,n2.
@@ -74,16 +80,16 @@ lemma lveq_fwd_bind_abst_length_le: ∀I1,I2,L1,L2,V2,n1,n2.
 (* Inversion lemmas with length for local environments **********************)
 
 lemma lveq_inv_void_dx_length: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2.ⓧ → |L1| ≤ |L2| →
-                               ∃∃m2. L1 ≋ ⓧ*[n1, m2] L2 & n2 = ⫯m2 & n1 ≤ m2.
+                               ∃∃m2. L1 ≋ ⓧ*[n1, m2] L2 & 0 = n1 & ⫯m2 = n2.
 #L1 #L2 #n1 #n2 #H #HL12
-lapply (lveq_fwd_length … H) normalize >plus_n_Sm #H0
+lapply (lveq_fwd_length_plus … H) normalize >plus_n_Sm #H0
 lapply (plus2_inv_le_sn … H0 HL12) -H0 -HL12 #H0
-elim (le_inv_S1 … H0) -H0 #m2 #Hm2 #H0 destruct
-/3 width=4 by lveq_inv_void_succ_dx, ex3_intro/
+elim (le_inv_S1 … H0) -H0 #m2 #_ #H0 destruct
+elim (lveq_inv_void_succ_dx … H) -H /2 width=3 by ex3_intro/
 qed-.
 
 lemma lveq_inv_void_sn_length: ∀L1,L2,n1,n2. L1.ⓧ ≋ⓧ*[n1, n2] L2 → |L2| ≤ |L1| →
-                               ∃∃m1. L1 ≋ ⓧ*[m1, n2] L2 & n1 = ⫯m1 & n2 ≤ m1.
+                               ∃∃m1. L1 ≋ ⓧ*[m1, n2] L2 & ⫯m1 = n1 & 0 = n2.
 #L1 #L2 #n1 #n2 #H #HL
 lapply (lveq_sym … H) -H #H
 elim (lveq_inv_void_dx_length … H HL) -H -HL
index 3f64c33b02d7ee4b753a04bcea7a060b5770dce5..3d3417dba8897d7d6f9e834d9f295404ec5cb865 100644 (file)
@@ -16,131 +16,38 @@ include "basic_2/syntax/lveq_length.ma".
 
 (* EQUIVALENCE FOR LOCAL ENVIRONMENTS UP TO EXCLUSION BINDERS ***************)
 
-(* Main forward lemmas ******************************************************)
-
-theorem lveq_fwd_inj_succ_zero: ∀L1,L2,n1. L1 ≋ⓧ*[⫯n1, 0] L2 →
-                                ∀m1,m2. L1 ≋ⓧ*[m1, m2] L2 → ∃x1. ⫯x1 = m1.
-#L1 #L2 #n1 #Hn #m1 #m2 #Hm
-lapply (lveq_fwd_length … Hn) -Hn <plus_n_O #Hn
-lapply (lveq_fwd_length … Hm) -Hm >Hn >associative_plus -Hn #Hm
-lapply (injective_plus_r … Hm) -Hm
-<plus_S1 /2 width=2 by ex_intro/
-qed-.
-
-theorem lveq_fwd_inj_zero_succ: ∀L1,L2,n2. L1 ≋ⓧ*[0, ⫯n2] L2 →
-                                ∀m1,m2. L1 ≋ⓧ*[m1, m2] L2 → ∃x2. ⫯x2 = m2.
-/4 width=6 by lveq_fwd_inj_succ_zero, lveq_sym/ qed-. (* auto: 2x lveq_sym *)
-
-theorem lveq_fwd_inj_succ_sn: ∀L1,L2,n1,n2. L1 ≋ⓧ*[⫯n1, n2] L2 →
-                              ∀m1,m2. L1 ≋ⓧ*[m1, m2] L2 →
-                              ∨∨ ∃x. ⫯x = n2 | ∃x. ⫯x = m1.
-#L1 #L2 #n1 * [2: #n2 ] /3 width=2 by ex_intro, or_introl/
-#Hn #m1 #m2 #Hm @or_intror @lveq_fwd_inj_succ_zero /width=6 by/ (**) (* auto fails *)
-qed-.
-
-theorem lveq_fwd_inj_succ_dx: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, ⫯n2] L2 →
-                              ∀m1,m2. L1 ≋ⓧ*[m1, m2] L2 →
-                              ∨∨ ∃x. ⫯x = n1 | ∃x. ⫯x = m2.
-/4 width=6 by lveq_fwd_inj_succ_sn, lveq_sym/ qed-. (* auto: 2x lveq_sym *) 
-
 (* Main inversion lemmas ****************************************************)
 
-theorem lveq_inv_pair_sn: ∀K1,K2,n. K1 ≋ⓧ*[n, n] K2 →
-                          ∀I1,I2,V,m1,m2. K1.ⓑ{I1}V ≋ⓧ*[m1, m2] K2.ⓘ{I2} →
-                          ∧∧ 0 = m1 & 0 = m2.
-#K1 #K2 #n #HK #I1 #I2 #V #m1 #m2 #H
+theorem lveq_inv_bind: ∀K1,K2. K1 ≋ⓧ*[0, 0] K2 →
+                       ∀I1,I2,m1,m2. K1.ⓘ{I1} ≋ⓧ*[m1, m2] K2.ⓘ{I2} →
+                       ∧∧ 0 = m1 & 0 = m2.
+#K1 #K2 #HK #I1 #I2 #m1 #m2 #H
 lapply (lveq_fwd_length_eq … HK) -HK #HK
-lapply (lveq_fwd_pair_sn … H) #H0 destruct
-<(lveq_inj_length … H) -H normalize /3 width=1 by conj, eq_f/
+elim (lveq_inj_length … H) -H normalize /3 width=1 by conj, eq_f/
 qed-.
 
-theorem lveq_inv_pair_dx: ∀K1,K2,n. K1 ≋ⓧ*[n, n] K2 →
-                          ∀I1,I2,V,m1,m2. K1.ⓘ{I1} ≋ⓧ*[m1, m2] K2.ⓑ{I2}V →
-                          ∧∧ 0 = m1 & 0 = m2.
-/4 width=8 by lveq_inv_pair_sn, lveq_sym, commutative_and/ qed-.
-
 theorem lveq_inj: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
                   ∀m1,m2. L1 ≋ⓧ*[m1, m2] L2 →
                   ∧∧ n1 = m1 & n2 = m2.
-#L1 #L2 @(f2_ind ?? length2 ?? L1 L2) -L1 -L2
-#x #IH * [2: #L1 #I1 ] * [2,4: #L2 #I2 ]
-[ cases I1 -I1 [ * | #I1 #V1 ] cases I2 -I2 [1,3: * |*: #I2 #V2 ] ]
-#Hx #n1 #n2 #Hn #m1 #m2 #Hm destruct
-[ elim (lveq_fwd_void_void … Hn) * #x #H destruct
-  elim (lveq_fwd_void_void … Hm) * #y #H destruct
-  [ lapply (lveq_inv_void_succ_sn … Hn) -Hn #Hn
-    lapply (lveq_inv_void_succ_sn … Hm) -Hm #Hm
-    elim (IH … Hn … Hm) -IH -Hn -Hm // #H1 #H2 destruct
-    /2 width=1 by conj/
-  | elim (lveq_fwd_inj_succ_sn … Hn … Hm) * #z #H destruct
-    [ lapply (lveq_inv_void_succ_dx … Hn) -Hn #Hn
-      lapply (lveq_inv_void_succ_dx … Hm) -Hm #Hm
-      elim (IH … Hn … Hm) -IH -Hn -Hm [2: normalize // ] #H1 #H2 destruct (**) (* avoid normalize *)
-      /2 width=1 by conj/
-    | lapply (lveq_inv_void_succ_sn … Hn) -Hn #Hn
-      lapply (lveq_inv_void_succ_sn … Hm) -Hm #Hm
-      elim (IH … Hn … Hm) -IH -Hn -Hm // #H1 #H2 destruct
-      /2 width=1 by conj/
-    ]
-  | elim (lveq_fwd_inj_succ_dx … Hn … Hm) * #z #H destruct
-    [ lapply (lveq_inv_void_succ_sn … Hn) -Hn #Hn
-      lapply (lveq_inv_void_succ_sn … Hm) -Hm #Hm
-      elim (IH … Hn … Hm) -IH -Hn -Hm // #H1 #H2 destruct
-      /2 width=1 by conj/
-    | lapply (lveq_inv_void_succ_dx … Hn) -Hn #Hn
-      lapply (lveq_inv_void_succ_dx … Hm) -Hm #Hm
-      elim (IH … Hn … Hm) -IH -Hn -Hm [2: normalize // ] #H1 #H2 destruct (**) (* avoid normalize *)
-      /2 width=1 by conj/
-    ]
-  | lapply (lveq_inv_void_succ_dx … Hn) -Hn #Hn
-    lapply (lveq_inv_void_succ_dx … Hm) -Hm #Hm
-    elim (IH … Hn … Hm) -IH -Hn -Hm [2: normalize // ] #H1 #H2 destruct (**) (* avoid normalize *)
-    /2 width=1 by conj/
-  ]
-| lapply (lveq_fwd_abst_bind_length_le … Hn) #HL
-  elim (le_to_or_lt_eq … HL) -HL #HL
-  [ elim (lveq_inv_void_dx_length … Hn) -Hn // #x1 #Hn #H #_ destruct
-    elim (lveq_inv_void_dx_length … Hm) -Hm // #y1 #Hm #H #_ destruct
-    elim (IH … Hn … Hm) -IH -Hn -Hm -HL [2: normalize // ] #H1 #H2 destruct (**) (* avoid normalize *)
-    /2 width=1 by conj/
-  | elim (lveq_eq_ex … HL) -HL #x #HL
-    elim (lveq_inv_pair_sn … HL … Hn) -Hn #H1 #H2 destruct
-    elim (lveq_inv_pair_sn … HL … Hm) -Hm #H1 #H2 destruct
-    /2 width=1 by conj/
-  ]
-| lapply (lveq_fwd_bind_abst_length_le … Hn) #HL
-  elim (le_to_or_lt_eq … HL) -HL #HL
-  [ elim (lveq_inv_void_sn_length … Hn) -Hn // #x1 #Hn #H #_ destruct
-    elim (lveq_inv_void_sn_length … Hm) -Hm // #y1 #Hm #H #_ destruct
-    elim (IH … Hn … Hm) -IH -Hn -Hm -HL // #H1 #H2 destruct
-    /2 width=1 by conj/
-  | lapply (sym_eq ??? HL) -HL #HL
-    elim (lveq_eq_ex … HL) -HL #x #HL
-    elim (lveq_inv_pair_dx … HL … Hn) -Hn #H1 #H2 destruct
-    elim (lveq_inv_pair_dx … HL … Hm) -Hm #H1 #H2 destruct
-    /2 width=1 by conj/
-  ]
-| elim (lveq_inv_pair_pair… Hn) -Hn #x #_ #H1 #H2 destruct
-  elim (lveq_inv_pair_pair… Hm) -Hm #y #_ #H1 #H2 destruct
-  /2 width=1 by conj/
-| elim (lveq_inv_atom_bind … Hn) -Hn #x #Hn #H1 #H2 destruct
-  elim (lveq_inv_atom_bind … Hm) -Hm #y #Hm #H1 #H2 destruct
-  elim (IH … Hn … Hm) -IH -Hn -Hm /2 width=1 by conj/
-| elim (lveq_inv_bind_atom … Hn) -Hn #x #Hn #H1 #H2 destruct
-  elim (lveq_inv_bind_atom … Hm) -Hm #y #Hm #H1 #H2 destruct
-  elim (IH … Hn … Hm) -IH -Hn -Hm /2 width=1 by conj/
-| elim (lveq_inv_atom_atom … Hn) -Hn #H1 #H2 destruct
-  elim (lveq_inv_atom_atom … Hm) -Hm #H1 #H2 destruct
-  /2 width=1 by conj/
-]
+#L1 #L2 #n1 #n2 #Hn #m1 #m2 #Hm
+elim (lveq_fwd_length … Hn) -Hn #H1 #H2 destruct
+elim (lveq_fwd_length … Hm) -Hm #H1 #H2 destruct
+/2 width=1 by conj/
 qed-.
 
-theorem lveq_inj_void_sn: ∀K1,K2,n1,n2. K1 ≋ⓧ*[n1, n2] K2 →
-                          ∀m1,m2. K1.ⓧ ≋ⓧ*[m1, m2] K2 →
-                          ∧∧ ⫯n1 = m1 & n2 = m2.
-/3 width=4 by lveq_inj, lveq_void_sn/ qed-.
+theorem lveq_inj_void_sn_ge: ∀K1,K2. |K2| ≤ |K1| →
+                             ∀n1,n2. K1 ≋ⓧ*[n1, n2] K2 →
+                             ∀m1,m2. K1.ⓧ ≋ⓧ*[m1, m2] K2 →
+                             ∧∧ ⫯n1 = m1 & 0 = m2 & 0 = n2.
+#L1 #L2 #HL #n1 #n2 #Hn #m1 #m2 #Hm
+elim (lveq_fwd_length … Hn) -Hn #H1 #H2 destruct
+elim (lveq_fwd_length … Hm) -Hm #H1 #H2 destruct
+>length_bind >eq_minus_S_pred >(eq_minus_O … HL)
+/3 width=4 by plus_minus, and3_intro/
+qed-.
 
-theorem lveq_inj_void_dx: ∀K1,K2,n1,n2. K1 ≋ⓧ*[n1, n2] K2 →
-                          ∀m1,m2. K1 ≋ⓧ*[m1, m2] K2.ⓧ →
-                          ∧∧ n1 = m1 & ⫯n2 = m2.
-/3 width=4 by lveq_inj, lveq_void_dx/ qed-.
+theorem lveq_inj_void_dx_le: ∀K1,K2. |K1| ≤ |K2| →
+                             ∀n1,n2. K1 ≋ⓧ*[n1, n2] K2 →
+                             ∀m1,m2. K1 ≋ⓧ*[m1, m2] K2.ⓧ →
+                             ∧∧ ⫯n2 = m2 & 0 = m1 & 0 = n1.
+/3 width=5 by lveq_inj_void_sn_ge, lveq_sym/ qed-. (* auto: 2x lveq_sym *)
index 0dd04cb2cc7f415862dd4cf324671f220b7eb7ac..618d4550e67992ebbda9b034faa1b4be84ca8b8a 100644 (file)
@@ -239,16 +239,15 @@ table {
    class "red"
    [ { "syntax" * } {
         [ { "equivalence up to exclusion binders" * } {
-             [ [ "" ] "voids ( ⓧ*[?]? ≋ ⓧ*[?]? )" "voids_length" + "voids_voids" * ]
+             [ [ "for lenvs" ] "lveq ( ? ≋ⓧ*[?,?] ? )" "lveq_length" + "lveq_lveq" * ]
           }
         ]
-   
-        [ { "append for local environments" * } {
-             [ [ "" ] "append ( ? @@ ? )" "append_length" * ]
+        [ { "append" * } {
+             [ [ "for lenvs" ] "append ( ? @@ ? )" "append_length" * ]
           }
         ]
-        [ { "head equivalence for terms" * } {
-             [ [ "" ] "theq ( ? ⩳[?,?] ? )" "theq_simple" + "theq_tdeq" + "theq_theq" + "theq_simple_vector" * ]
+        [ { "head equivalence" * } {
+             [ [ "for terms" ] "theq ( ? ⩳[?,?] ? )" "theq_simple" + "theq_tdeq" + "theq_theq" + "theq_simple_vector" * ]
           }
         ]
         [ { "degree-based equivalence" * } {