]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma
renaming in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / lex_tc.ma
index 9e48f4d361c825462cb63b51ffdcc70bf6c5876c..00b9db820a136e35b75a6355c56392eaf2f47677 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "basic_2/syntax/ext2_tc.ma".
-include "basic_2/relocation/lexs_tc.ma".
+include "basic_2/relocation/sex_tc.ma".
 include "basic_2/relocation/lex.ma".
 
 alias symbol "subseteq" = "relation inclusion".
@@ -23,14 +23,14 @@ alias symbol "subseteq" = "relation inclusion".
 (* Inversion lemmas with transitive closure *********************************)
 
 (* Basic_2A1: was: lpx_sn_LTC_TC_lpx_sn *)
-lemma lex_inv_CTC: ∀R. c_reflexive … R →
-                   lex (CTC … R) ⊆ TC … (lex R).
+lemma lex_inv_CTC (R): c_reflexive … R →
+                       lex (CTC … R) ⊆ TC … (lex R).
 #R #HR #L1 #L2 *
-/5 width=11 by lexs_inv_tc_dx, lexs_co, ext2_inv_tc, ext2_refl, monotonic_TC, ex2_intro/
+/5 width=11 by sex_inv_tc_dx, sex_co, ext2_inv_tc, ext2_refl, monotonic_TC, ex2_intro/
 qed-.
 
-lemma s_rs_transitive_lex_inv_isid: ∀R. s_rs_transitive … R (λ_.lex R) →
-                                    s_rs_transitive_isid cfull (cext2 R).
+lemma s_rs_transitive_lex_inv_isid (R): s_rs_transitive … R (λ_.lex R) →
+                                        s_rs_transitive_isid cfull (cext2 R).
 #R #HR #f #Hf #L2 #T1 #T2 #H #L1 #HL12
 elim (ext2_tc … H) -H
 [ /3 width=1 by ext2_inv_tc, ext2_unit/
@@ -43,21 +43,47 @@ qed-.
 (* Properties with transitive closure ***************************************)
 
 (* Basic_2A1: was: TC_lpx_sn_inv_lpx_sn_LTC *)
-lemma lex_CTC: ∀R. s_rs_transitive … R (λ_. lex R) →
-               TC … (lex R) ⊆ lex (CTC … R).
+lemma lex_CTC (R): s_rs_transitive … R (λ_. lex R) →
+                   TC … (lex R) ⊆ lex (CTC … R).
 #R #HR #L1 #L2 #HL12
-lapply (monotonic_TC … (lexs cfull (cext2 R) 𝐈𝐝) … HL12) -HL12
-[ #L1 #L2 * /3 width=3 by lexs_eq_repl_fwd, eq_id_inv_isid/
-| /5 width=9 by s_rs_transitive_lex_inv_isid, lexs_tc_dx, lexs_co, ext2_tc, ex2_intro/
+lapply (monotonic_TC … (sex cfull (cext2 R) 𝐈𝐝) … HL12) -HL12
+[ #L1 #L2 * /3 width=3 by sex_eq_repl_fwd, eq_id_inv_isid/
+| /5 width=9 by s_rs_transitive_lex_inv_isid, sex_tc_dx, sex_co, ext2_tc, ex2_intro/
 ]
 qed-.
 
-lemma lex_CTC_step_dx: ∀R. c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
-                       ∀L1,L. lex (CTC … R) L1 L →
-                       ∀L2. lex R L L2 → lex (CTC … R) L1 L2.
+lemma lex_CTC_inj (R): s_rs_transitive … R (λ_. lex R) →
+                       (lex R) ⊆ lex (CTC … R).
+/3 width=1 by lex_CTC, inj/ qed-.
+
+lemma lex_CTC_step_dx (R): c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
+                           ∀L1,L. lex (CTC … R) L1 L →
+                           ∀L2. lex R L L2 → lex (CTC … R) L1 L2.
 /4 width=3 by lex_CTC, lex_inv_CTC, step/ qed-.
 
-lemma lex_CTC_step_sn: ∀R. c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
-                       ∀L1,L. lex R L1 L →
-                       ∀L2. lex (CTC … R) L L2 → lex (CTC … R) L1 L2.
+lemma lex_CTC_step_sn (R): c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
+                           ∀L1,L. lex R L1 L →
+                           ∀L2. lex (CTC … R) L L2 → lex (CTC … R) L1 L2.
 /4 width=3 by lex_CTC, lex_inv_CTC, TC_strap/ qed-.
+
+(* Eliminators with transitive closure **************************************)
+
+lemma lex_CTC_ind_sn (R) (L2): c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
+                               ∀Q:predicate lenv. Q L2 →
+                               (∀L1,L. L1 ⪤[R] L → L ⪤[CTC … R] L2 → Q L → Q L1) →
+                               ∀L1. L1 ⪤[CTC … R] L2 → Q L1.
+#R #L2 #H1R #H2R #Q #IH1 #IH2 #L1 #H
+lapply (lex_inv_CTC … H1R … H) -H #H
+@(TC_star_ind_dx ???????? H) -H
+/3 width=4 by lex_CTC, lex_refl/
+qed-.
+
+lemma lex_CTC_ind_dx (R) (L1): c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
+                               ∀Q:predicate lenv. Q L1 →
+                               (∀L,L2. L1 ⪤[CTC … R] L → L ⪤[R] L2 → Q L → Q L2) →
+                               ∀L2. L1 ⪤[CTC … R] L2 → Q L2.
+#R #L1 #H1R #H2R #Q #IH1 #IH2 #L2 #H
+lapply (lex_inv_CTC … H1R … H) -H #H
+@(TC_star_ind ???????? H) -H
+/3 width=4 by lex_CTC, lex_refl/
+qed-.