]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/gcp.ma
slight refactoring in the proof of strong normalization
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / gcp.ma
index 67dd3f6d65c8eda631e9d458af5ee88936f16478..3cab7cbaf6de724ddb9d4001bf8be4ce35748616 100644 (file)
@@ -17,38 +17,42 @@ include "basic_2/multiple/drops.ma".
 
 (* GENERIC COMPUTATION PROPERTIES *******************************************)
 
+definition nf ≝ λRR:relation4 genv lenv term term. λRS:relation term.
+                λG,L,T. NF … (RR G L) RS T.
+
 definition candidate: Type[0] ≝ relation3 genv lenv term.
 
 definition CP0 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
-                 ∀G,L0,L,T,T0,s,d,e. NF … (RR G L) RS T →
-                 ⇩[s, d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → NF … (RR G L0) RS T0.
-
-definition CP0s ≝ λRR:relation4 genv lenv term term. λRS:relation term.
-                  ∀G,L0,L,s,des. ⇩*[s, des] L0 ≡ L →
-                  ∀T,T0. ⇧*[des] T ≡ T0 →
-                  NF … (RR G L) RS T → NF … (RR G L0) RS T0.
+                 ∀G. l_liftable1 (nf RR RS G) (Ⓕ).
 
 definition CP1 ≝ λRR:relation4 genv lenv term term. λRS:relation term.
                  ∀G,L. ∃k. NF … (RR G L) RS (⋆k).
 
-definition CP2 ≝ λRP:candidate.
+definition CP2 ≝ λRP:candidate. ∀G. l_liftable1 (RP G) (Ⓕ).
+
+definition CP3 ≝ λRP:candidate.
                  ∀G,L,T,k. RP G L (ⓐ⋆k.T) → RP G L T.
 
 (* requirements for generic computation properties *)
 record gcp (RR:relation4 genv lenv term term) (RS:relation term) (RP:candidate) : Prop ≝
 { cp0: CP0 RR RS;
   cp1: CP1 RR RS;
-  cp2: CP2 RP
+  cp2: CP2 RP;
+  cp3: CP3 RP
 }.
 
 (* Basic properties *********************************************************)
 
 (* Basic_1: was: nf2_lift1 *)
-lemma gcp_lifts: ∀RR,RS. CP0 RR RS → CP0s RR RS.
-#RR #RS #HRR #G #L1 #L2 #s #des #H elim H -L1 -L2 -des
-[ #L #T1 #T2 #H #HT1
-  <(lifts_inv_nil … H) -H //
-| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #T2 #T1 #H #HLT2
-  elim (lifts_inv_cons … H) -H /3 width=10 by/
-]
+lemma gcp0_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀G. l_liftables1 (nf RR RS G) (Ⓕ).
+#RR #RS #RP #H #G @l1_liftable_liftables @(cp0 … H)
+qed.
+
+lemma gcp2_lifts: ∀RR,RS,RP. gcp RR RS RP → ∀G. l_liftables1 (RP G) (Ⓕ).
+#RR #RS #RP #H #G @l1_liftable_liftables @(cp2 … H)
+qed.
+
+(* Basic_1: was only: sns3_lifts1 *)
+lemma gcp2_lifts_all: ∀RR,RS,RP. gcp RR RS RP → ∀G. l_liftables1_all (RP G) (Ⓕ).
+#RR #RS #RP #H #G @l1_liftables_liftables_all /2 width=7 by gcp2_lifts/
 qed.