]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/conversion/cpc.ma
- first prroperties on lfsx proved
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / conversion / cpc.ma
index 1268d8887ab0a535da4271d0111bb7ff4d6c7894..156374e3fe96ef719d3750bc2cb5e8f4d9f6d3f7 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/notation/relations/pconv_4.ma".
-include "basic_2/reduction/cpr.ma".
+include "basic_2/notation/relations/pconv_5.ma".
+include "basic_2/rt_transition/cpm.ma".
 
-(* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************)
+(* CONTEXT-SENSITIVE PARALLEL R-CONVERSION FOR TERMS ************************)
 
-definition cpc: relation4 genv lenv term term ≝
-                λG,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 ∨ ⦃G, L⦄ ⊢ T2 ➡ T1.
+definition cpc: sh → relation4 genv lenv term term ≝
+                λh,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h] T2 ∨ ⦃G, L⦄ ⊢ T2 ➡[h] T1.
 
 interpretation
-   "context-sensitive parallel conversion (term)"
-   'PConv G L T1 T2 = (cpc G L T1 T2).
+   "context-sensitive parallel r-conversion (term)"
+   'PConv h G L T1 T2 = (cpc h G L T1 T2).
 
 (* Basic properties *********************************************************)
 
-lemma cpc_refl: ∀G,L. reflexive … (cpc G L).
+lemma cpc_refl: ∀h,G,L. reflexive … (cpc h G L).
 /2 width=1 by or_intror/ qed.
 
-lemma cpc_sym: ∀G,L. symmetric … (cpc L G).
-#G #L #T1 #T2 * /2 width=1 by or_introl, or_intror/
+lemma cpc_sym: ∀h,G,L. symmetric … (cpc h L G).
+#h #G #L #T1 #T2 * /2 width=1 by or_introl, or_intror/
 qed-.
 
 (* Basic forward lemmas *****************************************************)
 
-lemma cpc_fwd_cpr: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌ T2 → ∃∃T. ⦃G, L⦄ ⊢ T1 ➡ T & ⦃G, L⦄ ⊢ T2 ➡ T.
-#G #L #T1 #T2 * /2 width=3 by ex2_intro/
+lemma cpc_fwd_cpr: ∀h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌[h] T2 →
+                   ∃∃T. ⦃G, L⦄ ⊢ T1 ➡[h] T & ⦃G, L⦄ ⊢ T2 ➡[h] T.
+#h #G #L #T1 #T2 * /2 width=3 by ex2_intro/
 qed-.