]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/static/rex_fsle.ma
more additions and corrections for the article
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / rex_fsle.ma
index 1a2829da3c91dde8c3f87e124ee8749686c940c9..7974bb900c8a1233f1139bd1cfd2a4e8f5a083c2 100644 (file)
@@ -30,10 +30,11 @@ definition rex_fsle_compatible: predicate (relation3 …) ≝ λRN.
 
 (* Basic inversions with free variables inclusion for restricted closures ***)
 
-lemma frees_sex_conf: ∀R. rex_fsge_compatible R →
-                      ∀L1,T,f1. L1 ⊢ 𝐅*⦃T⦄ ≘ f1 →
-                      ∀L2. L1 ⪤[cext2 R,cfull,f1] L2 →
-                      ∃∃f2. L2 ⊢ 𝐅*⦃T⦄ ≘ f2 & f2 ⊆ f1.
+lemma frees_sex_conf (R):
+      rex_fsge_compatible R →
+      ∀L1,T,f1. L1 ⊢ 𝐅+⦃T⦄ ≘ f1 →
+      ∀L2. L1 ⪤[cext2 R,cfull,f1] L2 →
+      ∃∃f2. L2 ⊢ 𝐅+⦃T⦄ ≘ f2 & f2 ⊆ f1.
 #R #HR #L1 #T #f1 #Hf1 #L2 #H1L
 lapply (HR L1 L2 T ?) /2 width=3 by ex2_intro/ #H2L
 @(fsle_frees_trans_eq … H2L … Hf1) /3 width=4 by sex_fwd_length, sym_eq/
@@ -42,26 +43,29 @@ qed-.
 (* 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 fsge_rex_trans: ∀R,L1,T1,T2. ⦃L1,T1⦄ ⊆ ⦃L1,T2⦄ →
-                      ∀L2. L1 ⪤[R,T2] L2 → L1 ⪤[R,T1] L2.
+lemma fsge_rex_trans (R):
+      ∀L1,T1,T2. ⦃L1,T1⦄ ⊆ ⦃L1,T2⦄ →
+      ∀L2. L1 ⪤[R,T2] L2 → L1 ⪤[R,T1] L2.
 #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 rex_inv_frees, sle_sex_trans, ex2_intro/
 qed-.
 
-lemma rex_sym: ∀R. rex_fsge_compatible R →
-               (∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) →
-               ∀T. symmetric … (rex R T).
+lemma rex_sym (R):
+      rex_fsge_compatible R →
+      (∀L1,L2,T1,T2. R L1 T1 T2 → R L2 T2 T1) →
+      ∀T. symmetric … (rex R T).
 #R #H1R #H2R #T #L1 #L2
 * #f1 #Hf1 #HL12
 elim (frees_sex_conf … Hf1 … HL12) -Hf1 //
 /5 width=5 by sle_sex_trans, sex_sym, cext2_sym, ex2_intro/
 qed-.
 
-lemma rex_pair_sn_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
-                         rex_fsge_compatible R1 →
-                         ∀L1,L2,V. L1 ⪤[R1,V] L2 → ∀I,T.
-                         ∃∃L. L1 ⪤[R1,②{I}V.T] L & L ⪤[R2,V] L2.
+lemma rex_pair_sn_split (R1) (R2):
+      (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
+      rex_fsge_compatible R1 →
+      ∀L1,L2,V. L1 ⪤[R1,V] L2 → ∀I,T.
+      ∃∃L. L1 ⪤[R1,②{I}V.T] L & L ⪤[R2,V] L2.
 #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #V * #f #Hf #HL12 * [ #p ] #I #T
 [ elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg
   elim (frees_inv_bind … Hg) #y1 #y2 #H #_ #Hy
@@ -77,10 +81,11 @@ elim (frees_sex_conf … Hf … H) -Hf -H
 /4 width=7 by sle_sex_trans, ex2_intro/
 qed-.
 
-lemma rex_flat_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
-                         rex_fsge_compatible R1 →
-                         ∀L1,L2,T. L1 ⪤[R1,T] L2 → ∀I,V.
-                         ∃∃L. L1 ⪤[R1,ⓕ{I}V.T] L & L ⪤[R2,T] L2.
+lemma rex_flat_dx_split (R1) (R2):
+      (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
+      rex_fsge_compatible R1 →
+      ∀L1,L2,T. L1 ⪤[R1,T] L2 → ∀I,V.
+      ∃∃L. L1 ⪤[R1,ⓕ{I}V.T] L & L ⪤[R2,T] L2.
 #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #I #V
 elim (frees_total L1 (ⓕ{I}V.T)) #g #Hg
 elim (frees_inv_flat … Hg) #y1 #y2 #_ #H #Hy
@@ -93,10 +98,11 @@ elim (frees_sex_conf … Hf … H) -Hf -H
 /4 width=7 by sle_sex_trans, ex2_intro/
 qed-.
 
-lemma rex_bind_dx_split: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
-                         rex_fsge_compatible R1 →
-                         ∀I,L1,L2,V1,T. L1.ⓑ{I}V1 ⪤[R1,T] L2 → ∀p.
-                         ∃∃L,V. L1 ⪤[R1,ⓑ{p,I}V1.T] L & L.ⓑ{I}V ⪤[R2,T] L2 & R1 L1 V1 V.
+lemma rex_bind_dx_split (R1) (R2):
+      (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
+      rex_fsge_compatible R1 →
+      ∀I,L1,L2,V1,T. L1.ⓑ{I}V1 ⪤[R1,T] L2 → ∀p.
+      ∃∃L,V. L1 ⪤[R1,ⓑ{p,I}V1.T] L & L.ⓑ{I}V ⪤[R2,T] L2 & R1 L1 V1 V.
 #R1 #R2 #HR1 #HR2 #HR #I #L1 #L2 #V1 #T * #f #Hf #HL12 #p
 elim (frees_total L1 (ⓑ{p,I}V1.T)) #g #Hg
 elim (frees_inv_bind … Hg) #y1 #y2 #_ #H #Hy
@@ -113,10 +119,11 @@ elim (frees_sex_conf … Hf … H0) -Hf -H0
 /4 width=7 by sle_sex_trans, ex3_2_intro, ex2_intro/
 qed-.
 
-lemma rex_bind_dx_split_void: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
-                              rex_fsge_compatible R1 →
-                              ∀L1,L2,T. L1.ⓧ ⪤[R1,T] L2 → ∀p,I,V.
-                              ∃∃L. L1 ⪤[R1,ⓑ{p,I}V.T] L & L.ⓧ ⪤[R2,T] L2.
+lemma rex_bind_dx_split_void (R1) (R2):
+      (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) →
+      rex_fsge_compatible R1 →
+      ∀L1,L2,T. L1.ⓧ ⪤[R1,T] L2 → ∀p,I,V.
+      ∃∃L. L1 ⪤[R1,ⓑ{p,I}V.T] L & L.ⓧ ⪤[R2,T] L2.
 #R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #p #I #V
 elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg
 elim (frees_inv_bind_void … Hg) #y1 #y2 #_ #H #Hy
@@ -135,11 +142,10 @@ qed-.
 
 (* Main properties with free variables inclusion for restricted closures ****)
 
-theorem rex_conf: ∀R1,R2.
-                  rex_fsge_compatible R1 →
-                  rex_fsge_compatible R2 →
-                  R_confluent2_rex R1 R2 R1 R2 →
-                  ∀T. confluent2 … (rex R1 T) (rex R2 T).
+theorem rex_conf (R1) (R2):
+        rex_fsge_compatible R1 → rex_fsge_compatible R2 →
+        R_confluent2_rex R1 R2 R1 R2 →
+        ∀T. confluent2 … (rex R1 T) (rex R2 T).
 #R1 #R2 #HR1 #HR2 #HR12 #T #L0 #L1 * #f1 #Hf1 #HL01 #L2 * #f #Hf #HL02
 lapply (frees_mono … Hf1 … Hf) -Hf1 #Hf12
 lapply (sex_eq_repl_back … HL01 … Hf12) -f1 #HL01
@@ -164,10 +170,9 @@ elim (sex_conf … HL01 … HL02) /2 width=3 by ex2_intro/ [ | -HL01 -HL02 ]
 ]
 qed-.
 
-theorem rex_trans_fsle: ∀R1,R2,R3.
-                        rex_fsle_compatible R1 → f_transitive_next R1 R2 R3 →
-                        ∀L1,L,T. L1 ⪤[R1,T] L →
-                        ∀L2. L ⪤[R2,T] L2 → L1 ⪤[R3,T] L2.
+theorem rex_trans_fsle (R1) (R2) (R3):
+        rex_fsle_compatible R1 → f_transitive_next R1 R2 R3 →
+        ∀L1,L,T. L1 ⪤[R1,T] L → ∀L2. L ⪤[R2,T] L2 → L1 ⪤[R3,T] L2.
 #R1 #R2 #R3 #H1R #H2R #L1 #L #T #H
 lapply (H1R … H) -H1R #H0
 cases H -H #f1 #Hf1 #HL1 #L2 * #f2 #Hf2 #HL2