]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_bind.ma
renaming in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / lifts_bind.ma
index 65a8e5f22582a919aec39c5f5dd43407e2862bf5..fd56f4e6975b0a5a823b5321adc9d9182042d7dc 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/syntax/bind.ma".
+include "basic_2/syntax/ext2.ma".
 include "basic_2/relocation/lifts.ma".
 
 (* GENERIC RELOCATION FOR BINDERS *******************************************)
@@ -28,37 +28,37 @@ interpretation "generic relocation (binder for local environments)"
 
 (* Basic_inversion lemmas **************************************************)
 
-lemma liftsb_inv_unit_sn: â\88\80f,I,Z2. â¬\86*[f] BUnit I â\89¡ Z2 → Z2 = BUnit I.
+lemma liftsb_inv_unit_sn: â\88\80f,I,Z2. â¬\86*[f] BUnit I â\89\98 Z2 → Z2 = BUnit I.
 /2 width=2 by ext2_inv_unit_sn/ qed-.
 
-lemma liftsb_inv_pair_sn: â\88\80f:rtmap. â\88\80Z2,I,V1. â¬\86*[f] BPair I V1 â\89¡ Z2 →
-                          â\88\83â\88\83V2. â¬\86*[f] V1 â\89¡ V2 & Z2 = BPair I V2.
+lemma liftsb_inv_pair_sn: â\88\80f:rtmap. â\88\80Z2,I,V1. â¬\86*[f] BPair I V1 â\89\98 Z2 →
+                          â\88\83â\88\83V2. â¬\86*[f] V1 â\89\98 V2 & Z2 = BPair I V2.
 /2 width=1 by ext2_inv_pair_sn/ qed-.
 
-lemma liftsb_inv_unit_dx: â\88\80f,I,Z1. â¬\86*[f] Z1 â\89¡ BUnit I → Z1 = BUnit I.
+lemma liftsb_inv_unit_dx: â\88\80f,I,Z1. â¬\86*[f] Z1 â\89\98 BUnit I → Z1 = BUnit I.
 /2 width=2 by ext2_inv_unit_dx/ qed-.
 
-lemma liftsb_inv_pair_dx: â\88\80f:rtmap. â\88\80Z1,I,V2. â¬\86*[f] Z1 â\89¡ BPair I V2 →
-                          â\88\83â\88\83V1. â¬\86*[f] V1 â\89¡ V2 & Z1 = BPair I V1.
+lemma liftsb_inv_pair_dx: â\88\80f:rtmap. â\88\80Z1,I,V2. â¬\86*[f] Z1 â\89\98 BPair I V2 →
+                          â\88\83â\88\83V1. â¬\86*[f] V1 â\89\98 V2 & Z1 = BPair I V1.
 /2 width=1 by ext2_inv_pair_dx/ qed-.
 
 (* Basic properties *********************************************************)
 
-lemma liftsb_eq_repl_back: â\88\80I1,I2. eq_repl_back â\80¦ (λf. â¬\86*[f] I1 â\89¡ I2).
+lemma liftsb_eq_repl_back: â\88\80I1,I2. eq_repl_back â\80¦ (λf. â¬\86*[f] I1 â\89\98 I2).
 #I1 #I2 #f1 * -I1 -I2 /3 width=3 by lifts_eq_repl_back, ext2_pair/
 qed-.
 
 lemma liftsb_refl: ∀f. 𝐈⦃f⦄ → reflexive … (liftsb f).
 /3 width=1 by lifts_refl, ext2_refl/ qed.
 
-lemma liftsb_total: â\88\80I1,f. â\88\83I2. â¬\86*[f] I1 â\89¡ I2.
+lemma liftsb_total: â\88\80I1,f. â\88\83I2. â¬\86*[f] I1 â\89\98 I2.
 * [2: #I #T1 #f elim (lifts_total T1 f) ]
 /3 width=2 by ext2_unit, ext2_pair, ex_intro/
 qed-.
 
-lemma liftsb_split_trans: â\88\80f,I1,I2. â¬\86*[f] I1 â\89¡ I2 →
-                          â\88\80f1,f2. f2 â\8a\9a f1 â\89¡ f →
-                          â\88\83â\88\83I. â¬\86*[f1] I1 â\89¡ I & â¬\86*[f2] I â\89¡ I2.
+lemma liftsb_split_trans: â\88\80f,I1,I2. â¬\86*[f] I1 â\89\98 I2 →
+                          â\88\80f1,f2. f2 â\8a\9a f1 â\89\98 f →
+                          â\88\83â\88\83I. â¬\86*[f1] I1 â\89\98 I & â¬\86*[f2] I â\89\98 I2.
 #f #I1 #I2 * -I1 -I2 /2 width=3 by ext2_unit, ex2_intro/
 #I #V1 #V2 #HV12 #f1 #f2 #Hf elim (lifts_split_trans … HV12 … Hf) -f
 /3 width=3 by ext2_pair, ex2_intro/
@@ -66,6 +66,6 @@ qed-.
 
 (* Basic forward lemmas *****************************************************)
 
-lemma liftsb_fwd_isid: â\88\80f,I1,I2. â¬\86*[f] I1 â\89¡ I2 → 𝐈⦃f⦄ → I1 = I2.
+lemma liftsb_fwd_isid: â\88\80f,I1,I2. â¬\86*[f] I1 â\89\98 I2 → 𝐈⦃f⦄ → I1 = I2.
 #f #I1 #I2 * -I1 -I2 /3 width=3 by lifts_fwd_isid, eq_f2/
 qed-.