]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sle.ma
- exclusion binder added in local environments
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_sle.ma
index c5d9428a562f07cb48c74c9e29d275627458a570..120902ce5598229299c5cc650a56667b6117b2e9 100644 (file)
@@ -13,6 +13,7 @@
 (**************************************************************************)
 
 include "ground_2/relocation/rtmap_isid.ma".
+include "ground_2/relocation/rtmap_isdiv.ma".
 
 (* RELOCATION MAP ***********************************************************)
 
@@ -27,11 +28,26 @@ interpretation "inclusion (rtmap)"
 
 (* Basic properties *********************************************************)
 
+axiom sle_eq_repl_back1: ∀f2. eq_repl_back … (λf1. f1 ⊆ f2).
+
+lemma sle_eq_repl_fwd1: ∀f2. eq_repl_fwd … (λf1. f1 ⊆ f2).
+#f2 @eq_repl_sym /2 width=3 by sle_eq_repl_back1/
+qed-.
+
+axiom sle_eq_repl_back2: ∀f1. eq_repl_back … (λf2. f1 ⊆ f2).
+
+lemma sle_eq_repl_fwd2: ∀f1. eq_repl_fwd … (λf2. f1 ⊆ f2).
+#f1 @eq_repl_sym /2 width=3 by sle_eq_repl_back2/
+qed-.
+
 corec lemma sle_refl: ∀f. f ⊆ f.
 #f cases (pn_split f) * #g #H
 [ @(sle_push … H H) | @(sle_next … H H) ] -H //
 qed.
 
+lemma sle_refl_eq: ∀f1,f2. f1 ≗ f2 → f1 ⊆ f2.
+/2 width=3 by sle_eq_repl_back2/ qed.
+
 (* Basic inversion lemmas ***************************************************)
 
 lemma sle_inv_xp: ∀g1,g2. g1 ⊆ g2 → ∀f2. ↑f2 = g2 →
@@ -147,3 +163,21 @@ corec lemma sle_inv_isid_dx: ∀f1,f2. f1 ⊆ f2 → 𝐈⦃f2⦄ → 𝐈⦃f1
 lapply (isid_inv_push … H ??) -H
 /3 width=3 by isid_push/
 qed-.
+
+(* Properties with isdiv ****************************************************)
+
+corec lemma sle_isdiv_dx: ∀f2. 𝛀⦃f2⦄ → ∀f1. f1 ⊆ f2.
+#f2 * -f2
+#f2 #g2 #Hf2 #H2 #f1 cases (pn_split f1) *
+/3 width=5 by sle_weak, sle_next/
+qed.
+
+(* Inversion lemmas with isdiv **********************************************)
+
+corec lemma sle_inv_isdiv_sn: ∀f1,f2. f1 ⊆ f2 → 𝛀⦃f1⦄ → 𝛀⦃f2⦄.
+#f1 #f2 * -f1 -f2
+#f1 #f2 #g1 #g2 #Hf * * #H
+[1,3: elim (isdiv_inv_push … H) // ]
+lapply (isdiv_inv_next … H ??) -H
+/3 width=3 by isdiv_next/
+qed-.