]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpms_cpms.ma
milestone uupdate in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpms_cpms.ma
index 5fe790c896c8e86dcafaf9c14a7cdae38eeba385..fcd97f3b905b5192dc3c7cc908bfebe0ef7bab2a 100644 (file)
@@ -12,6 +12,8 @@
 (*                                                                        *)
 (**************************************************************************)
 
+include "ground_2/xoa/ex_5_7.ma".
+include "basic_2/rt_transition/cpm_lsubr.ma".
 include "basic_2/rt_computation/cpms_drops.ma".
 include "basic_2/rt_computation/cprs.ma".
 
@@ -19,19 +21,6 @@ include "basic_2/rt_computation/cprs.ma".
 
 (* Main properties **********************************************************)
 
-(* Basic_2A1: uses: lstas_scpds_trans scpds_strap2 *)
-theorem cpms_trans (h) (G) (L):
-                   ∀n1,T1,T. ⦃G, L⦄ ⊢ T1 ➡*[n1, h] T →
-                   ∀n2,T2. ⦃G, L⦄ ⊢ T ➡*[n2, h] T2 → ⦃G, L⦄ ⊢ T1 ➡*[n1+n2, h] T2.
-/2 width=3 by ltc_trans/ qed-.
-
-(* Basic_2A1: uses: scpds_cprs_trans *)
-theorem cpms_cprs_trans (n) (h) (G) (L):
-                        ∀T1,T. ⦃G, L⦄ ⊢ T1 ➡*[n, h] T →
-                        ∀T2. ⦃G, L⦄ ⊢ T ➡*[h] T2 → ⦃G, L⦄ ⊢ T1 ➡*[n, h] T2.
-#n #h #G #L #T1 #T #HT1 #T2 #HT2 >(plus_n_O … n)
-/2 width=3 by cpms_trans/ qed-.
-
 (* Basic_2A1: includes: cprs_bind *)
 theorem cpms_bind (n) (h) (G) (L):
                   ∀I,V1,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ➡*[n, h] T2 →
@@ -55,28 +44,6 @@ theorem cpms_appl (n) (h) (G) (L):
 ]
 qed.
 
-lemma cpms_inv_plus (h) (G) (L): ∀n1,n2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[n1+n2, h] T2 →
-                                 ∃∃T. ⦃G, L⦄ ⊢ T1 ➡*[n1, h] T & ⦃G, L⦄ ⊢ T ➡*[n2, h] T2.
-#h #G #L #n1 elim n1 -n1 /2 width=3 by ex2_intro/
-#n1 #IH #n2 #T1 #T2 <plus_S1 #H
-elim (cpms_inv_succ_sn … H) -H #T0 #HT10 #HT02
-elim (IH … HT02) -IH -HT02 #T #HT0 #HT2
-lapply (cpms_trans … HT10 … HT0) -T0 #HT1
-/2 width=3 by ex2_intro/
-qed-.
-
-lemma cpms_cast (n) (h) (G) (L):
-                ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[n, h] T2 →
-                ∀U1,U2. ⦃G, L⦄ ⊢ U1 ➡*[n, h] U2 →
-                ⦃G, L⦄ ⊢ ⓝU1.T1 ➡*[n, h] ⓝU2.T2.
-#n #h #G #L #T1 #T2 #H @(cpms_ind_sn … H) -T1 -n
-[ /3 width=3 by cpms_cast_sn/
-| #n1 #n2 #T1 #T #HT1 #_ #IH #U1 #U2 #H
-  elim (cpms_inv_plus … H) -H #U #HU1 #HU2
-  /3 width=3 by cpms_trans, cpms_cast_sn/
-]
-qed.
-
 (* Basic_2A1: includes: cprs_beta_rc *)
 theorem cpms_beta_rc (n) (h) (G) (L):
                      ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h] V2 →
@@ -128,71 +95,84 @@ theorem cpms_theta (n) (h) (G) (L):
   /3 width=3 by cpr_pair_sn, cpms_step_sn/
 ]
 qed.
-(*
-(* Advanced inversion lemmas ************************************************)
 
-(* Basic_1: was pr3_gen_appl *)
-lemma cprs_inv_appl1: ∀G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡* U2 →
-                      ∨∨ ∃∃V2,T2.       ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L⦄ ⊢ T1 ➡* T2 &
-                                        U2 = ⓐV2. T2
-                       | ∃∃a,W,T.       ⦃G, L⦄ ⊢ T1 ➡* ⓛ{a}W.T &
-                                        ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡* U2
-                       | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡* V0 & ⬆[0,1] V0 ≘ V2 &
-                                        ⦃G, L⦄ ⊢ T1 ➡* ⓓ{a}V.T &
-                                        ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡* U2.
-#G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
-#U #U2 #_ #HU2 * *
-[ #V0 #T0 #HV10 #HT10 #H destruct
-  elim (cpr_inv_appl1 … HU2) -HU2 *
-  [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5 by cprs_strap1, or3_intro0, ex3_2_intro/
-  | #a #V2 #W #W2 #T #T2 #HV02 #HW2 #HT2 #H1 #H2 destruct
-    lapply (cprs_strap1 … HV10 … HV02) -V0 #HV12
-    lapply (lsubr_cpr_trans … HT2 (L.ⓓⓝW.V1) ?) -HT2
-    /5 width=5 by cprs_bind, cprs_flat_dx, cpr_cprs, lsubr_beta, ex2_3_intro, or3_intro1/
-  | #a #V #V2 #W0 #W2 #T #T2 #HV0 #HV2 #HW02 #HT2 #H1 #H2 destruct
-    /5 width=10 by cprs_flat_sn, cprs_bind_dx, cprs_strap1, ex4_5_intro, or3_intro2/
-  ]
-| /4 width=9 by cprs_strap1, or3_intro1, ex2_3_intro/
-| /4 width=11 by cprs_strap1, or3_intro2, ex4_5_intro/
-]
-qed-.
+(* Basic_2A1: uses: lstas_scpds_trans scpds_strap2 *)
+theorem cpms_trans (h) (G) (L):
+                   ∀n1,T1,T. ⦃G, L⦄ ⊢ T1 ➡*[n1, h] T →
+                   ∀n2,T2. ⦃G, L⦄ ⊢ T ➡*[n2, h] T2 → ⦃G, L⦄ ⊢ T1 ➡*[n1+n2, h] T2.
+/2 width=3 by ltc_trans/ qed-.
 
-(* Advanced inversion lemmas ************************************************)
+(* Basic_2A1: uses: scpds_cprs_trans *)
+theorem cpms_cprs_trans (n) (h) (G) (L):
+                        ∀T1,T. ⦃G, L⦄ ⊢ T1 ➡*[n, h] T →
+                        ∀T2. ⦃G, L⦄ ⊢ T ➡*[h] T2 → ⦃G, L⦄ ⊢ T1 ➡*[n, h] T2.
+#n #h #G #L #T1 #T #HT1 #T2 #HT2 >(plus_n_O … n)
+/2 width=3 by cpms_trans/ qed-.
 
-lemma scpds_inv_abst1: ∀h,o,a,G,L,V1,T1,U2,d. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, o, d] U2 →
-                       ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, o, d] T2 &
-                                U2 = ⓛ{a}V2.T2.
-#h #o #a #G #L #V1 #T1 #U2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
-lapply (da_inv_bind … Hd1) -Hd1 #Hd1
-elim (lstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct
-elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct
-/3 width=6 by ex4_2_intro, ex3_2_intro/
-qed-.
+(* Advanced inversion lemmas ************************************************)
 
-lemma scpds_inv_abbr_abst: ∀h,o,a1,a2,G,L,V1,W2,T1,T2,d. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, o, d] ⓛ{a2}W2.T2 →
-                           ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, o, d] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≘ T & a1 = true.
-#h #o #a1 #a2 #G #L #V1 #W2 #T1 #T2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2
-lapply (da_inv_bind … Hd1) -Hd1 #Hd1
-elim (lstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct
-elim (cprs_inv_abbr1 … H2) -H2 *
-[ #V2 #U2 #HV12 #HU12 #H destruct
-| /3 width=6 by ex4_2_intro, ex3_intro/
+lemma cpms_inv_appl_sn (n) (h) (G) (L):
+                       ∀V1,T1,X2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡*[n, h] X2 →
+                       ∨∨ ∃∃V2,T2.
+                            ⦃G, L⦄ ⊢ V1 ➡*[h] V2 & ⦃G, L⦄ ⊢ T1 ➡*[n, h] T2 &
+                            X2 = ⓐV2.T2
+                        | ∃∃n1,n2,p,W,T.
+                            ⦃G, L⦄ ⊢ T1 ➡*[n1, h] ⓛ{p}W.T & ⦃G, L⦄ ⊢ ⓓ{p}ⓝW.V1.T ➡*[n2, h] X2 &
+                            n1 + n2 = n
+                        | ∃∃n1,n2,p,V0,V2,V,T.
+                            ⦃G, L⦄ ⊢ V1 ➡*[h] V0 & ⬆*[1] V0 ≘ V2 &
+                            ⦃G, L⦄ ⊢ T1 ➡*[n1, h] ⓓ{p}V.T & ⦃G, L⦄ ⊢ ⓓ{p}V.ⓐV2.T ➡*[n2, h] X2 &
+                            n1 + n2 = n.
+#n #h #G #L #V1 #T1 #U2 #H
+@(cpms_ind_dx … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/
+#n1 #n2 #U #U2 #_ * *
+[ #V0 #T0 #HV10 #HT10 #H #HU2 destruct
+  elim (cpm_inv_appl1 … HU2) -HU2 *
+  [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5 by cpms_step_dx, or3_intro0, ex3_2_intro/
+  | #p #V2 #W #W2 #T #T2 #HV02 #HW2 #HT2 #H1 #H2 destruct
+    lapply (cprs_step_dx … HV10 … HV02) -V0 #HV12
+    lapply (lsubr_cpm_trans … HT2 (L.ⓓⓝW.V1) ?) -HT2
+    /5 width=8 by cprs_flat_dx, cpms_bind, cpm_cpms, lsubr_beta, ex3_5_intro, or3_intro1/
+  | #p #V #V2 #W0 #W2 #T #T2 #HV0 #HV2 #HW02 #HT2 #H1 #H2 destruct
+    /6 width=12 by cprs_step_dx, cpm_cpms, cpm_appl, cpm_bind, ex5_7_intro, or3_intro2/
+  ]
+| #m1 #m2 #p #W #T #HT1 #HTU #H #HU2 destruct
+  lapply (cpms_step_dx … HTU … HU2) -U #H
+  @or3_intro1 @(ex3_5_intro … HT1 H) // (**) (* auto fails *)
+| #m1 #m2 #p #V2 #W2 #V #T #HV12 #HVW2 #HT1 #HTU #H #HU2 destruct
+  lapply (cpms_step_dx … HTU … HU2) -U #H
+  @or3_intro2 @(ex5_7_intro … HV12 HVW2 HT1 H) // (**) (* auto fails *)
 ]
 qed-.
 
-lemma scpds_inv_lstas_eq: ∀h,o,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*➡*[h, o, d] T2 →
-                          ∀T. ⦃G, L⦄ ⊢ T1 •*[h, d] T → ⦃G, L⦄ ⊢ T ➡* T2.
-#h #o #G #L #T1 #T2 #d2 *
-#T0 #d1 #_ #_ #HT10 #HT02 #T #HT1
-lapply (lstas_mono … HT10 … HT1) #H destruct //
+lemma cpms_inv_plus (h) (G) (L): ∀n1,n2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[n1+n2, h] T2 →
+                                 ∃∃T. ⦃G, L⦄ ⊢ T1 ➡*[n1, h] T & ⦃G, L⦄ ⊢ T ➡*[n2, h] T2.
+#h #G #L #n1 elim n1 -n1 /2 width=3 by ex2_intro/
+#n1 #IH #n2 #T1 #T2 <plus_S1 #H
+elim (cpms_inv_succ_sn … H) -H #T0 #HT10 #HT02
+elim (IH … HT02) -IH -HT02 #T #HT0 #HT2
+lapply (cpms_trans … HT10 … HT0) -T0 #HT1
+/2 width=3 by ex2_intro/
 qed-.
 
-(* Main properties **********************************************************)
+(* Advanced main properties *************************************************)
+
+theorem cpms_cast (n) (h) (G) (L):
+                  ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[n, h] T2 →
+                  ∀U1,U2. ⦃G, L⦄ ⊢ U1 ➡*[n, h] U2 →
+                  ⦃G, L⦄ ⊢ ⓝU1.T1 ➡*[n, h] ⓝU2.T2.
+#n #h #G #L #T1 #T2 #H @(cpms_ind_sn … H) -T1 -n
+[ /3 width=3 by cpms_cast_sn/
+| #n1 #n2 #T1 #T #HT1 #_ #IH #U1 #U2 #H
+  elim (cpms_inv_plus … H) -H #U #HU1 #HU2
+  /3 width=3 by cpms_trans, cpms_cast_sn/
+]
+qed.
 
-theorem scpds_conf_eq: ∀h,o,G,L,T0,T1,d. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T1 →
-                       ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, o, d] T2 →
-                       ∃∃T. ⦃G, L⦄ ⊢ T1 ➡* T & ⦃G, L⦄ ⊢ T2 ➡* T.
-#h #o #G #L #T0 #T1 #d0 * #U1 #d1 #_ #_ #H1 #HUT1 #T2 * #U2 #d2 #_ #_ #H2 #HUT2 -d1 -d2
-lapply (lstas_mono … H1 … H2) #H destruct -h -d0 /2 width=3 by cprs_conf/
+theorem cpms_trans_swap (h) (G) (L) (T1):
+        ∀n1,T. ⦃G,L⦄ ⊢ T1 ➡*[n1,h] T → ∀n2,T2. ⦃G,L⦄ ⊢ T ➡*[n2,h] T2 →
+        ∃∃T0. ⦃G,L⦄ ⊢ T1 ➡*[n2,h] T0 & ⦃G,L⦄ ⊢ T0 ➡*[n1,h] T2.
+#h #G #L #T1 #n1 #T #HT1 #n2 #T2 #HT2
+lapply (cpms_trans … HT1 … HT2) -T <commutative_plus #HT12
+/2 width=1 by cpms_inv_plus/
 qed-.
-*)