]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sor.ma
more on lfpx_frees.ma ...
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_sor.ma
index 02bac68a04d5594e3daec27c9fbf52407d9239ab..f86e7f8092784bb8e8ddd8e07c07dc9134aba43e 100644 (file)
@@ -266,7 +266,47 @@ corec lemma sor_sym: ∀f1,f2,f. f1 ⋓ f2 ≡ f → f2 ⋓ f1 ≡ f.
 [ @sor_pp | @sor_pn | @sor_np | @sor_nn ] /2 width=7 by/
 qed-.
 
-(* Properies on test for identity *******************************************)
+(* Main properties **********************************************************)
+
+axiom monotonic_sle_sor: ∀f1,g1. f1 ⊆ g1 → ∀f2,g2. f2 ⊆ g2 →
+                         ∀f. f1 ⋓ f2 ≡ f → ∀g. g1 ⋓ g2 ≡ g → f ⊆ g.
+
+axiom sor_trans1: ∀f0,f3,f4. f0 ⋓ f3 ≡ f4 →
+                  ∀f1,f2. f1 ⋓ f2 ≡ f0 →
+                  ∀f. f2 ⋓ f3 ≡ f → f1 ⋓ f ≡ f4.
+
+axiom sor_trans2: ∀f1,f0,f4. f1 ⋓ f0 ≡ f4 →
+                  ∀f2, f3. f2 ⋓ f3 ≡ f0 →
+                  ∀f. f1 ⋓ f2 ≡ f → f ⋓ f3 ≡ f4.
+
+(* Properties with tail *****************************************************)
+
+lemma sor_tl: ∀f1,f2,f. f1 ⋓ f2 ≡ f → ⫱f1 ⋓ ⫱f2 ≡ ⫱f.
+#f1 cases (pn_split f1) * #g1 #H1
+#f2 cases (pn_split f2) * #g2 #H2
+#f #Hf
+[ cases (sor_inv_ppx … Hf … H1 H2)
+| cases (sor_inv_pnx … Hf … H1 H2)
+| cases (sor_inv_npx … Hf … H1 H2)
+| cases (sor_inv_nnx … Hf … H1 H2)
+] -Hf #g #Hg #H destruct //
+qed.
+
+lemma sor_xxn_tl: ∀g1,g2,g. g1 ⋓ g2 ≡ g → ∀f. ⫯f = g →
+                  (∃∃f1,f2. f1 ⋓ f2 ≡ f & ⫯f1 = g1 & ⫱g2 = f2) ∨
+                  (∃∃f1,f2. f1 ⋓ f2 ≡ f & ⫱g1 = f1 & ⫯f2 = g2).
+#g1 #g2 #g #H #f #H0 elim (sor_inv_xxn … H … H0) -H -H0 *
+/3 width=5 by ex3_2_intro, or_introl, or_intror/
+qed-.
+
+(* Properties with iterated tail ********************************************)
+
+lemma sor_tls: ∀f1,f2,f. f1 ⋓ f2 ≡ f →
+               ∀n. ⫱*[n]f1 ⋓ ⫱*[n]f2 ≡ ⫱*[n]f.
+#f1 #f2 #f #Hf #n elim n -n /2 width=1 by sor_tl/
+qed.
+
+(* Properies with test for identity *****************************************)
 
 corec lemma sor_isid_sn: ∀f1. 𝐈⦃f1⦄ → ∀f2. f1 ⋓ f2 ≡ f2.
 #f1 * -f1
@@ -283,7 +323,7 @@ qed.
 lemma sor_isid: ∀f1,f2,f. 𝐈⦃f1⦄ → 𝐈⦃f2⦄ → 𝐈⦃f⦄ → f1 ⋓ f2 ≡ f.
 /4 width=3 by sor_eq_repl_back2, sor_eq_repl_back1, isid_inv_eq_repl/ qed.
 
-(* Inversion lemmas on test for identity ************************************)
+(* Inversion lemmas with test for identity **********************************)
 
 lemma sor_isid_inv_sn: ∀f1,f2,f. f1 ⋓ f2 ≡ f → 𝐈⦃f1⦄ → f2 ≗ f.
 /3 width=4 by sor_isid_sn, sor_mono/
@@ -310,7 +350,7 @@ qed-.
 lemma sor_inv_isid3: ∀f1,f2,f. f1 ⋓ f2 ≡ f → 𝐈⦃f⦄ → 𝐈⦃f1⦄ ∧ 𝐈⦃f2⦄.
 /3 width=4 by sor_fwd_isid2, sor_fwd_isid1, conj/ qed-.
 
-(* Properties on finite colength assignment *********************************)
+(* Properties with finite colength assignment *******************************)
 
 lemma sor_fcla_ex: ∀f1,n1. 𝐂⦃f1⦄ ≡ n1 → ∀f2,n2. 𝐂⦃f2⦄ ≡ n2 →
                    ∃∃f,n. f1 ⋓ f2 ≡ f & 𝐂⦃f⦄ ≡ n & (n1 ∨ n2) ≤ n & n ≤ n1 + n2.
@@ -349,7 +389,7 @@ lemma sor_fwd_fcla_dx_ex: ∀f,n. 𝐂⦃f⦄ ≡ n → ∀f1,f2. f1 ⋓ f2 ≡
                           ∃∃n2.  𝐂⦃f2⦄ ≡ n2 & n2 ≤ n.
 /3 width=4 by sor_fwd_fcla_sn_ex, sor_sym/ qed-.
 
-(* Properties on test for finite colength ***********************************)
+(* Properties with test for finite colength *********************************)
 
 lemma sor_isfin_ex: ∀f1,f2. 𝐅⦃f1⦄ → 𝐅⦃f2⦄ → ∃∃f. f1 ⋓ f2 ≡ f & 𝐅⦃f⦄.
 #f1 #f2 * #n1 #H1 * #n2 #H2 elim (sor_fcla_ex … H1 … H2) -H1 -H2
@@ -378,7 +418,7 @@ qed-.
 lemma sor_inv_isfin3: ∀f1,f2,f. f1 ⋓ f2 ≡ f → 𝐅⦃f⦄ → 𝐅⦃f1⦄ ∧ 𝐅⦃f2⦄.
 /3 width=4 by sor_fwd_isfin_dx, sor_fwd_isfin_sn, conj/ qed-.
 
-(* Inversion lemmas on inclusion ********************************************)
+(* Inversion lemmas with inclusion ******************************************)
 
 corec lemma sor_inv_sle_sn: ∀f1,f2,f. f1 ⋓ f2 ≡ f → f1 ⊆ f.
 #f1 #f2 #f * -f1 -f2 -f
@@ -391,3 +431,11 @@ corec lemma sor_inv_sle_dx: ∀f1,f2,f. f1 ⋓ f2 ≡ f → f2 ⊆ f.
 #f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0
 /3 width=5 by sle_push, sle_next, sle_weak/
 qed-.
+
+(* Properties with inclusion ************************************************)
+
+lemma sor_sle_sn: ∀f1,f2,f. f1 ⋓ f2 ≡ f → ∀g. g ⊆ f1 → g ⊆ f.
+/3 width=4 by sor_inv_sle_sn, sle_trans/ qed.
+
+lemma sor_sle_dx: ∀f1,f2,f. f1 ⋓ f2 ≡ f → ∀g. g ⊆ f2 → g ⊆ f.
+/3 width=4 by sor_inv_sle_dx, sle_trans/ qed.