]> matita.cs.unibo.it Git - helm.git/commitdiff
the partial commit continues ...
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 5 Sep 2012 16:09:37 +0000 (16:09 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 5 Sep 2012 16:09:37 +0000 (16:09 +0000)
matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_append.ma
matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_cif.ma
matita/matita/contribs/lambda_delta/basic_2/reducibility/crf.ma
matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_append.ma
matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma
matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma
matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ltpss.ma
matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma

index 83b65a3353d47ac1f57f3d656cdd0d5338b00f55..2f693218856771dccd6bd96b156ef5d136b911ae 100644 (file)
@@ -29,6 +29,10 @@ lemma append_atom_sn: ∀L. ⋆ @@ L = L.
 #L elim L -L normalize //
 qed.
 
+lemma append_assoc: associative … append.
+#L1 #L2 #L3 elim L3 -L3 normalize //
+qed.
+
 lemma append_length: ∀L1,L2. |L1 @@ L2| = |L1| + |L2|.
 #L1 #L2 elim L2 -L2 normalize //
 qed.
index 33e52e1b72ccd97d73317b62edb5a0cc4875eeed..51ce95aa066356cb2f94e304e38727f00d6aa7b3 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/substitution/tps_lift.ma".
-include "basic_2/unfold/tpss.ma".
 include "basic_2/reducibility/cif.ma".
 include "basic_2/reducibility/cnf_lift.ma".
 
index 3823b4c29824a2c010b40ec0a51b80d6c07a6cd6..5015f033c4b1b7a20937e8ffb86a170758b0d793 100644 (file)
@@ -12,7 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/grammar/term_simple.ma".
 include "basic_2/substitution/ldrop.ma".
 
 (* CONTEXT-SENSITIVE REDUCIBLE TERMS ****************************************)
index 3f6b2c53816ea3ed1a0b636b9eb2ae70a9e1196d..359d39c80c4ea0a59c82a9809a79504730cb1a29 100644 (file)
@@ -12,7 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/grammar/lenv_append.ma".
 include "basic_2/substitution/ldrop.ma".
 
 (* DROPPING *****************************************************************)
@@ -31,6 +30,8 @@ lemma ldrop_O1_append_sn_le: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → e ≤ |L1| →
                              ∀L. ⇩[0, e] L @@ L1 ≡ L @@ L2.
 /2 width=3 by ldrop_O1_append_sn_le_aux/ qed.
 
+(* Inversion lemmas on append for local environments ************************)
+
 lemma ldrop_O1_inv_append1_ge: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K →
                                |L2| ≤ e → ⇩[0, e - |L2|] L1 ≡ K.
 #K #L1 #L2 elim L2 -L2 normalize //
index 7e35dca7ca5e31ae427ae32bbe2ab50c19fcec83..0bde244171dedc979a8d303b1785f43b45ed3ac2 100644 (file)
@@ -12,8 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/grammar/lenv_top.ma".
-include "basic_2/substitution/ldrop.ma".
+include "basic_2/substitution/ldrop_append.ma".
 
 (* PARALLEL SUBSTITUTION ON TERMS *******************************************)
 
@@ -145,6 +144,15 @@ lemma tps_split_down: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 →
 ]
 qed.
 
+lemma tps_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶ [d, e] T2 →
+                  ∀L. L @@ K ⊢ T1 ▶ [d, e] T2.
+#K #T1 #T2 #d #e #H elim H -K -T1 -T2 -d -e // /2 width=1/
+#K #K0 #V #W #i #d #e #Hdi #Hide #HK0 #HVW #L
+lapply (ldrop_fwd_ldrop2_length … HK0) #H
+@(tps_subst … (L@@K0) … HVW) // (**) (* /3/ does not work *)
+@(ldrop_O1_append_sn_le … HK0) /2 width=2/
+qed.
+
 (* Basic inversion lemmas ***************************************************)
 
 fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀I. T1 = ⓪{I} →
@@ -251,17 +259,6 @@ lemma tps_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → #{T1} ≤ #{T2}.
 /3 by monotonic_le_plus_l, le_plus/ (**) (* just /3 width=1/ is too slow *)
 qed-.
 
-lemma tps_fwd_shift1: ∀L1,L,T1,T,d,e. L ⊢ L1 @@ T1 ▶ [d, e] T →
-                      ∃∃L2,T2. L1 𝟙 L2 & T = L2 @@ T2.
-#L1 @(lenv_ind_dx … L1) -L1
-[ #L #T1 #T #d #e #_ @ex2_2_intro [3: // |4: // |1,2: skip ] (**) (* /2 width=4/ does not work *)
-| #I #L1 #V1 #IH #L #T1 #T #d #e >shift_append_assoc #H
-  elim (tps_inv_bind1 … H) -H #V2 #T2 #_ #HT12 #H destruct
-  elim (IH … HT12) -IH -L -T1 -d -e #L2 #T #HL12 #H destruct
-  @(ex2_2_intro … (⋆.ⓑ{I}V2@@L2)) /2 width=4/ /3 width=2/
-]
-qed-.
-
 (* Basic_1: removed theorems 25:
             subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt
             subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans
index 89b19ea4576c8505f730d3c3115ab38e429edcd6..dfd1a19f6bb1271d27e01e3c84378571ba10f1a0 100644 (file)
@@ -31,79 +31,6 @@ inductive ltpss: nat → nat → relation lenv ≝
 interpretation "parallel unfold (local environment)"
    'PSubstStar L1 d e L2 = (ltpss d e L1 L2).
 
-(* Basic properties *********************************************************)
-
-lemma ltpss_tps2: ∀L1,L2,I,V1,V2,e.
-                  L1 ▶* [0, e] L2 → L2 ⊢ V1 ▶ [0, e] V2 →
-                  L1. ⓑ{I} V1 ▶* [0, e + 1] L2. ⓑ{I} V2.
-/3 width=1/ qed.
-
-lemma ltpss_tps1: ∀L1,L2,I,V1,V2,d,e.
-                  L1 ▶* [d, e] L2 → L2 ⊢ V1 ▶ [d, e] V2 →
-                  L1. ⓑ{I} V1 ▶* [d + 1, e] L2. ⓑ{I} V2.
-/3 width=1/ qed.
-
-lemma ltpss_tpss2_lt: ∀L1,L2,I,V1,V2,e.
-                      L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶* [0, e - 1] V2 →
-                      0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2.
-#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He
->(plus_minus_m_m e 1) /2 width=1/
-qed.
-
-lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e.
-                      L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶* [d - 1, e] V2 →
-                      0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2.
-#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd
->(plus_minus_m_m d 1) /2 width=1/
-qed.
-
-lemma ltpss_tps2_lt: ∀L1,L2,I,V1,V2,e.
-                     L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶ [0, e - 1] V2 →
-                     0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2.
-/3 width=1/ qed.
-
-lemma ltpss_tps1_lt: ∀L1,L2,I,V1,V2,d,e.
-                     L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶ [d - 1, e] V2 →
-                     0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2.
-/3 width=1/ qed.
-
-(* Basic_1: was by definition: csubst1_refl *)
-lemma ltpss_refl: ∀L,d,e. L ▶* [d, e] L.
-#L elim L -L //
-#L #I #V #IHL * /2 width=1/ * /2 width=1/
-qed.
-
-lemma ltpss_weak: ∀L1,L2,d1,e1. L1 ▶* [d1, e1] L2 →
-                  ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ▶* [d2, e2] L2.
-#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 //
-[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2
-  lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2;
-  lapply (lt_to_le_to_lt 0 … Hde2) // #He2
-  lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/
-| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12
-  >plus_plus_comm_23 in Hde12; #Hde12
-  elim (le_to_or_lt_eq 0 d2 ?) // #H destruct
-  [ lapply (le_plus_to_minus_r … Hde12) -Hde12 <plus_minus // #Hde12
-    lapply (le_plus_to_minus … Hd21) -Hd21 #Hd21 /3 width=5/
-  | -Hd21 normalize in Hde12;
-    lapply (lt_to_le_to_lt 0 … Hde12) // #He2
-    lapply (le_plus_to_minus_r … Hde12) -Hde12 /3 width=5/
-  ]
-]
-qed.  
-
-lemma ltpss_weak_all: ∀L1,L2,d,e. L1 ▶* [d, e] L2 → L1 ▶* [0, |L2|] L2.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-// /3 width=2/ /3 width=3/
-qed.
-
-(* Basic forward lemmas *****************************************************)
-
-lemma ltpss_fwd_length: ∀L1,L2,d,e. L1 ▶* [d, e] L2 → |L1| = |L2|.
-#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
-normalize //
-qed-.
-
 (* Basic inversion lemmas ***************************************************)
 
 fact ltpss_inv_refl_O2_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → e = 0 → L1 = L2.
@@ -219,6 +146,135 @@ lemma ltpss_inv_tpss12: ∀L1,K2,I,V2,d,e. L1 ▶* [d, e] K2. ⓑ{I} V2 → 0 <
                                  L1 = K1. ⓑ{I} V1.
 /2 width=3/ qed-.
 
+(* Basic properties *********************************************************)
+
+lemma ltpss_tps2: ∀L1,L2,I,V1,V2,e.
+                  L1 ▶* [0, e] L2 → L2 ⊢ V1 ▶ [0, e] V2 →
+                  L1. ⓑ{I} V1 ▶* [0, e + 1] L2. ⓑ{I} V2.
+/3 width=1/ qed.
+
+lemma ltpss_tps1: ∀L1,L2,I,V1,V2,d,e.
+                  L1 ▶* [d, e] L2 → L2 ⊢ V1 ▶ [d, e] V2 →
+                  L1. ⓑ{I} V1 ▶* [d + 1, e] L2. ⓑ{I} V2.
+/3 width=1/ qed.
+
+lemma ltpss_tpss2_lt: ∀L1,L2,I,V1,V2,e.
+                      L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶* [0, e - 1] V2 →
+                      0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2.
+#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He
+>(plus_minus_m_m e 1) /2 width=1/
+qed.
+
+lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e.
+                      L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶* [d - 1, e] V2 →
+                      0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2.
+#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd
+>(plus_minus_m_m d 1) /2 width=1/
+qed.
+
+lemma ltpss_tps2_lt: ∀L1,L2,I,V1,V2,e.
+                     L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶ [0, e - 1] V2 →
+                     0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2.
+/3 width=1/ qed.
+
+lemma ltpss_tps1_lt: ∀L1,L2,I,V1,V2,d,e.
+                     L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶ [d - 1, e] V2 →
+                     0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2.
+/3 width=1/ qed.
+
+(* Basic_1: was by definition: csubst1_refl *)
+lemma ltpss_refl: ∀L,d,e. L ▶* [d, e] L.
+#L elim L -L //
+#L #I #V #IHL * /2 width=1/ * /2 width=1/
+qed.
+
+lemma ltpss_weak: ∀L1,L2,d1,e1. L1 ▶* [d1, e1] L2 →
+                  ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ▶* [d2, e2] L2.
+#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 //
+[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2
+  lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2;
+  lapply (lt_to_le_to_lt 0 … Hde2) // #He2
+  lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/
+| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12
+  >plus_plus_comm_23 in Hde12; #Hde12
+  elim (le_to_or_lt_eq 0 d2 ?) // #H destruct
+  [ lapply (le_plus_to_minus_r … Hde12) -Hde12 <plus_minus // #Hde12
+    lapply (le_plus_to_minus … Hd21) -Hd21 #Hd21 /3 width=5/
+  | -Hd21 normalize in Hde12;
+    lapply (lt_to_le_to_lt 0 … Hde12) // #He2
+    lapply (le_plus_to_minus_r … Hde12) -Hde12 /3 width=5/
+  ]
+]
+qed.
+
+lemma ltpss_weak_all: ∀L1,L2,d,e. L1 ▶* [d, e] L2 → L1 ▶* [0, |L2|] L2.
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
+// /3 width=2/ /3 width=3/
+qed.
+
+fact ltpss_append_le_aux: ∀K1,K2,d,x. K1 ▶* [d, x] K2 → x = |K1| - d →
+                          ∀L1,L2,e. L1 ▶* [0, e] L2 → d ≤ |K1| →
+                          L1 @@ K1 ▶* [d, x + e] L2 @@ K2.
+#K1 #K2 #d #x #H elim H -K1 -K2 -d -x
+[ #d #x #H1 #L1 #L2 #e #HL12 #H2 destruct
+  lapply (le_n_O_to_eq … H2) -H2 #H destruct //
+| #K #I #V <minus_n_O normalize <plus_n_Sm #H destruct
+| #K1 #K2 #I #V1 #V2 #x #_ #HV12 <minus_n_O #IHK12 <minus_n_O #H #L1 #L2 #e #HL12 #_
+  lapply (injective_plus_l … H) -H #H destruct >plus_plus_comm_23
+  /4 width=5 by ltpss_tpss2, tpss_append, tpss_weak, monotonic_le_plus_r/ (**) (* too slow without trace *)
+| #K1 #K2 #I #V1 #V2 #d #x #_ #HV12 #IHK12 normalize <minus_le_minus_minus_comm // <minus_plus_m_m #H1 #L1 #L2 #e #HL12 #H2 destruct
+  lapply (le_plus_to_le_r … H2) -H2 #Hd
+  /4 width=5 by ltpss_tpss1, tpss_append, tpss_weak, monotonic_le_plus_r/ (**) (* too slow without trace *)
+]
+qed-.
+
+lemma ltpss_append_le: ∀K1,K2,d. K1 ▶* [d, |K1| - d] K2 →
+                       ∀L1,L2,e. L1 ▶* [0, e] L2 → d ≤ |K1| →
+                       L1 @@ K1 ▶* [d, |K1| - d + e] L2 @@ K2.
+/2 width=1 by ltpss_append_le_aux/ qed.
+
+lemma ltpss_append_ge: ∀K1,K2,d,e. K1 ▶* [d, e] K2 →
+                       ∀L1,L2. L1 ▶* [d - |K1|, e] L2 → |K1| ≤ d →
+                       L1 @@ K1 ▶* [d, e] L2 @@ K2.
+#K1 #K2 #d #e #H elim H -K1 -K2 -d -e
+[ #d #e #L1 #L2 <minus_n_O //
+| #K #I #V #L1 #L2 #_ #H 
+  lapply (le_n_O_to_eq … H) -H normalize <plus_n_Sm #H destruct
+| #K1 #K2 #I #V1 #V2 #e #_ #_ #_ #L1 #L2 #_ #H
+  lapply (le_n_O_to_eq … H) -H normalize <plus_n_Sm #H destruct
+| #K1 #K2 #I #V1 #V2 #d #e #_ #HV12 #IHK12 #L1 #L2
+  normalize <minus_le_minus_minus_comm // <minus_plus_m_m #HL12 #H
+  lapply (le_plus_to_le_r … H) -H /3 width=1/
+]
+qed.
+
+(* Basic forward lemmas *****************************************************)
+
+lemma ltpss_fwd_length: ∀L1,L2,d,e. L1 ▶* [d, e] L2 → |L1| = |L2|.
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e
+normalize //
+qed-.
+(*
+lemma tps_fwd_shift1: ∀L1,L,T1,T,d,e. L ⊢ L1 @@ T1 ▶ [d, e] T →
+                      ∃∃L2,T2. L @@ L1 ▶* [d + |L1|, e] L @@ L2 & T = L2 @@ T2.
+#L1 @(lenv_ind_dx … L1) -L1
+[ #L #T1 #T #d #e #_ @ex2_2_intro [3: // |4: // |1,2: skip ] (**) (* /2 width=4/ does not work *)
+| #I #L1 #V1 #IH #L #T1 #T #d #e >shift_append_assoc #H
+  elim (tps_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct
+  elim (IH … HT12) -IH -T1 #L2 #T #HL12 #H destruct
+  <append_assoc >append_length <associative_plus
+  @(ex2_2_intro … (⋆.ⓑ{I}V2@@L2)) /2 width=4/ <append_assoc normalize
+  lapply (ltpss_tps1 L … I … HV12) -HV12 // #HV12 
+  @ltpss_append_ge /2/
+(*  
+  
+  
+   /3 width=5/
+*)
+]
+qed-.
+*)
+
 (* Basic_1: removed theorems 28:
             csubst0_clear_O csubst0_drop_lt csubst0_drop_gt csubst0_drop_eq
             csubst0_clear_O_back csubst0_clear_S csubst0_clear_trans
index ec5037e8638ac01aeee082e88820c80955cfb338..f866adc126cd41eb357b208098568ee8a483a81c 100644 (file)
@@ -133,3 +133,14 @@ lemma ltpss_conf: ∀L0,L1,d1,e1. L0 ▶* [d1, e1] L1 →
                   ∀L2,d2,e2. L0 ▶* [d2, e2] L2 →
                   ∃∃L. L1 ▶* [d2, e2] L & L2 ▶* [d1, e1] L.
 /2 width=7/ qed.
+
+theorem ltpss_trans_eq: ∀L1,L,d,e. L1 ▶* [d, e] L →
+                        ∀L2. L ▶* [d, e] L2 → L1 ▶* [d, e] L2.
+#L1 #L #d #e #H elim H -L1 -L -d -e //
+[ #L1 #L #I #V1 #V #e #_ #HV1 #IHL1 #X #H
+  elim (ltpss_inv_tpss21 … H ?) -H // <minus_plus_m_m #L2 #V2 #HL2 #HV2 #H destruct
+  elim (ltpss_tpss_conf … HV1 … HL2) -HV1 #V0 #HV10 #HV0
+  elim (tpss_conf_eq … HV2 … HV0) -V #V #HV2 #HV0
+  lapply (tpss_trans_eq … HV10 … HV0) -V0 #HV1
+  @ltpss_tpss2 /2 width=1/
+        
\ No newline at end of file
index dd2c219ff94ab8ece693e793359d85d599c08082..a48f52ade156674fb5fca5a6b5d94f0a49bb1a14 100644 (file)
@@ -104,6 +104,11 @@ lapply (tpss_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12
 lapply (tpss_weak_top … HT12) //
 qed.
 
+lemma tpss_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶* [d, e] T2 →
+                   ∀L. L @@ K ⊢ T1 ▶* [d, e] T2.
+#K #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // /3 width=3/
+qed.
+
 (* Basic inversion lemmas ***************************************************)
 
 (* Note: this can be derived from tpss_inv_atom1 *)