]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground/relocation/tr_pat.ma
update in ground
[helm.git] / matita / matita / contribs / lambdadelta / ground / relocation / tr_pat.ma
index 208d9e7c0b37c665340e2ee9ab27d6ef76917bde..73b6dca66b6b85d3de7157b919295d62a3591d20 100644 (file)
@@ -21,18 +21,18 @@ include "ground/relocation/tr_map.ma".
 (* Constructions with pr_pat ***********************************************)
 
 (*** at_O1 *)
-lemma tr_pat_unit_sn: ∀i2,f. @❨𝟏,𝐭❨i2⨮f❩❩ ≘ i2.
+lemma tr_pat_unit_sn: ∀i2,f. @⧣❨𝟏,𝐭❨i2⨮f❩❩ ≘ i2.
 #i2 elim i2 -i2 /2 width=5 by pr_pat_refl, pr_pat_next/
 qed.
 
 (*** at_S1 *)
-lemma tr_pat_succ_sn: ∀p,f,i1,i2. @❨i1, 𝐭❨f❩❩ ≘ i2 → @❨↑i1, 𝐭❨p⨮f❩❩ ≘ i2+p.
+lemma tr_pat_succ_sn: ∀p,f,i1,i2. @⧣❨i1, 𝐭❨f❩❩ ≘ i2 → @⧣❨↑i1, 𝐭❨p⨮f❩❩ ≘ i2+p.
 #p elim p -p /3 width=7 by pr_pat_push, pr_pat_next/
 qed.
 
 (*** at_plus2 *)
 lemma tr_pat_plus_dx (f):
-      ∀i1,i,p,q. @❨i1, 𝐭❨p⨮f❩❩ ≘ i → @❨i1, 𝐭❨(p+q)⨮f❩❩ ≘ i+q.
+      ∀i1,i,p,q. @⧣❨i1, 𝐭❨p⨮f❩❩ ≘ i → @⧣❨i1, 𝐭❨(p+q)⨮f❩❩ ≘ i+q.
 #f #i1 #i #p #q #H elim q -q
 /2 width=5 by pr_pat_next/
 qed.
@@ -41,7 +41,7 @@ qed.
 
 (*** at_inv_O1 *)
 lemma tr_pat_inv_unit_sn (f):
-      ∀p,i2. @❨𝟏, 𝐭❨p⨮f❩❩ ≘ i2 → p = i2.
+      ∀p,i2. @⧣❨𝟏, 𝐭❨p⨮f❩❩ ≘ i2 → p = i2.
 #f #p elim p -p /2 width=6 by pr_pat_inv_unit_push/
 #p #IH #i2 #H elim (pr_pat_inv_next … H) -H [|*: // ]
 #j2 #Hj * -i2 /3 width=1 by eq_f/
@@ -49,8 +49,8 @@ qed-.
 
 (*** at_inv_S1 *)
 lemma tr_pat_inv_succ_sn (f):
-      ∀p,j1,i2. @❨↑j1, 𝐭❨p⨮f❩❩ ≘ i2 →
-      ∃∃j2. @❨j1, 𝐭❨f❩❩ ≘ j2 & j2+p = i2.
+      ∀p,j1,i2. @⧣❨↑j1, 𝐭❨p⨮f❩❩ ≘ i2 →
+      ∃∃j2. @⧣❨j1, 𝐭❨f❩❩ ≘ j2 & j2+p = i2.
 #f #p elim p -p /2 width=5 by pr_pat_inv_succ_push/
 #p #IH #j1 #i2 #H elim (pr_pat_inv_next … H) -H [|*: // ]
 #j2 #Hj * -i2 elim (IH … Hj) -IH -Hj
@@ -62,7 +62,7 @@ qed-.
 (* Note: a better conclusion would be: "i1 + ↓p ≤ i2" *)
 (*** at_increasing_plus *)
 lemma tr_pat_increasing_plus (f):
-      ∀p,i1,i2. @❨i1, 𝐭❨p⨮f❩❩ ≘ i2 → i1 + p ≤ ↑i2.
+      ∀p,i1,i2. @⧣❨i1, 𝐭❨p⨮f❩❩ ≘ i2 → i1 + p ≤ ↑i2.
 #f #p *
 [ #i2 #H <(tr_pat_inv_unit_sn … H) -i2 //
 | #i1 #i2 #H elim (tr_pat_inv_succ_sn … H) -H
@@ -73,7 +73,7 @@ qed-.
 
 (*** at_fwd_id *)
 lemma tr_pat_des_id (f):
-      ∀p,i. @❨i, 𝐭❨p⨮f❩❩ ≘ i → 𝟏 = p.
+      ∀p,i. @⧣❨i, 𝐭❨p⨮f❩❩ ≘ i → 𝟏 = p.
 #f #p #i #H lapply (pr_pat_des_id … H) -H #H
 elim (eq_inv_pr_push_cons … H) -H //
 qed-.