]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/computation/acp_cr.ma
- "functional" component moved to Apps_2
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / computation / acp_cr.ma
index 9f94d8969092f22419e62df7916a5e15fb378423..6eb71054b54a2ffc560c20f1cf6548245ab5a4ef 100644 (file)
@@ -27,7 +27,7 @@ definition S1 ≝ λRP,C:lenv→predicate term.
 (* Note: this is Tait's iii, or Girard's CR4 *)
 definition S2 ≝ λRR:lenv→relation term. λRS:relation term. λRP,C:lenv→predicate term.
                 ∀L,Vs. all … (RP L) Vs →
-                â\88\80T. ð\9d\95\8a[T] → NF … (RR L) RS T → C L (ⒶVs.T).
+                â\88\80T. ð\9d\90\92[T] → NF … (RR L) RS T → C L (ⒶVs.T).
 
 (* Note: this is Tait's ii *)
 definition S3 ≝ λRP,C:lenv→predicate term.
@@ -44,8 +44,8 @@ definition S5 ≝ λRP,C:lenv→predicate term.
 definition S6 ≝ λRP,C:lenv→predicate term.
                 ∀L,Vs,T,W. C L (ⒶVs. T) → RP L W → C L (ⒶVs. ⓣW. T).
 
-definition S7 ≝ λC:lenv→predicate term. ∀L1,L2,T1,T2,d,e.
-                C L1 T1 → ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → C L2 T2.
+definition S7 ≝ λC:lenv→predicate term. ∀L2,L1,T1,d,e.
+                C L1 T1 â\86\92 â\88\80T2. â\87©[d, e] L2 â\89¡ L1 â\86\92 â\87§[d, e] T1 â\89¡ T2 â\86\92 C L2 T2.
 
 definition S7s ≝ λC:lenv→predicate term.
                  ∀L1,L2,des. ⇩*[des] L2 ≡ L1 →
@@ -76,6 +76,7 @@ interpretation
 
 (* Basic properties *********************************************************)
 
+(* Basic_1: was: sc3_lift1 *)
 lemma acr_lifts: ∀C. S7 C → S7s C.
 #C #HC #L1 #L2 #des #H elim H -L1 -L2 -des
 [ #L #T1 #T2 #H #HT1
@@ -93,14 +94,18 @@ lemma rp_lifts: ∀RR,RS,RP. acr RR RS RP (λL,T. RP L T) →
 @(s7 … HRP)
 qed.
 
+(* Basic_1: was only: sns3_lifts1 *)
 lemma rp_liftsv_all: ∀RR,RS,RP. acr RR RS RP (λL,T. RP L T) →
-                     ∀des,L0,L,Vs,V0s. ⇧*[des] Vs ≡ V0s →  ⇩*[des] L0 ≡ L →
+                     ∀des,L0,L,Vs,V0s. ⇧*[des] Vs ≡ V0s → ⇩*[des] L0 ≡ L →
                      all … (RP L) Vs → all … (RP L0) V0s.
 #RR #RS #RP #HRP #des #L0 #L #Vs #V0s #H elim H -Vs -V0s normalize //
 #T1s #T2s #T1 #T2 #HT12 #_ #IHT2s #HL0 * #HT1 #HT1s
 @conj /2 width=1/ /2 width=6 by rp_lifts/
 qed.
 
+(* Basic_1: was: 
+   sc3_sn3 sc3_abst sc3_appl sc3_abbr sc3_bind sc3_cast sc3_lift
+*) 
 lemma aacr_acr: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
                 ∀A. acr RR RS RP (aacr RP A).
 #RR #RS #RP #H1RP #H2RP #A elim A -A normalize //
@@ -127,7 +132,7 @@ lemma aacr_acr: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
   >(at_mono … Hi1 … Hi0) in HL02; -i1 #HL02
   elim (ldrops_inv_skip2 … Hdes0 … H) -H -des0 #L2 #W1 #des0 #Hdes0 #HLK #HVW1 #H destruct
   elim (lift_total W1 0 (i0 + 1)) #W2 #HW12
-  elim (lifts_lift_trans … HVW1 … HW12 … Hdes0) // -Hdes0 -Hi0 #V3 #HV13 #HVW2
+  elim (lifts_lift_trans  … Hdes0 … HVW1 … HW12) // -Hdes0 -Hi0 #V3 #HV13 #HVW2
   >(lift_mono … HV13 … HV12) in HVW2; -V3 #HVW2
   @(s4 … IHA … (V0 :: V0s) … HW12 HL02) /3 width=4/
 | #L #V1s #V2s #HV12s #V #T #HA #HV #L0 #V10 #X #des #HB #HL0 #H
@@ -164,3 +169,6 @@ lapply (s1 … HCB) -HCB #HCB
 @(s3 … HCA … ◊) /2 width=6 by rp_lifts/
 @(s5 … HCA … ◊ ◊) // /2 width=1/ /2 width=3/
 qed.
+
+(* Basic_1: removed theorems 2: sc3_arity_gen sc3_repl *)
+(* Basic_1: removed local theorems 1: sc3_sn3_abst *)