]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma
- advances towards strong normalization
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cpm_drops.ma
index fbcd7958d4776e6ec180603c5e5601a0981d7e9b..790970db3cb5cec03c3e9debc49fb24f4219eda7 100644 (file)
@@ -78,18 +78,24 @@ qed-.
 
 (* Basic_1: includes: pr0_lift pr2_lift *)
 (* Basic_2A1: includes: cpr_lift *)
-lemma cpm_lifts: ∀n,h,G. d_liftable2 (cpm n h G).
+lemma cpm_lifts_sn: ∀n,h,G. d_liftable2_sn (cpm n h G).
 #n #h #G #K #T1 #T2 * #c #Hc #HT12 #b #f #L #HLK #U1 #HTU1
-elim (cpg_lifts … HT12 … HLK … HTU1) -K -T1
+elim (cpg_lifts_sn … HT12 … HLK … HTU1) -K -T1
 /3 width=5 by ex2_intro/
 qed-.
 
+lemma cpm_lifts_bi: ∀n,h,G. d_liftable2_bi (cpm n h G).
+/3 width=9 by cpm_lifts_sn, d_liftable2_sn_bi/ qed-.
+
 (* Inversion lemmas with generic slicing for local environments *************)
 
 (* Basic_1: includes: pr0_gen_lift pr2_gen_lift *)
 (* Basic_2A1: includes: cpr_inv_lift1 *)
-lemma cpm_inv_lifts1: ∀n,h,G. d_deliftable2_sn (cpm n h G).
+lemma cpm_inv_lifts_sn: ∀n,h,G. d_deliftable2_sn (cpm n h G).
 #n #h #G #L #U1 #U2 * #c #Hc #HU12 #b #f #K #HLK #T1 #HTU1
-elim (cpg_inv_lifts1 … HU12 … HLK … HTU1) -L -U1
+elim (cpg_inv_lifts_sn … HU12 … HLK … HTU1) -L -U1
 /3 width=5 by ex2_intro/
 qed-.
+
+lemma cpm_inv_lifts_bi: ∀n,h,G. d_deliftable2_bi (cpm n h G).
+/3 width=9 by cpm_inv_lifts_sn, d_deliftable2_sn_bi/ qed-.