]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_sor.ma
basic_2: stronger supclosure allows better inversion lemmas
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / relocation / rtmap_sor.ma
index b93214beadafbf4198d5d64d27bda5c9009ede38..640f9e265b41726ace787548be926c48227700f9 100644 (file)
@@ -72,6 +72,141 @@ try elim (discr_push_next … Hx2) try elim (discr_next_push … Hx2)
 /2 width=3 by ex2_intro/
 qed-.
 
+(* Advanced inversion lemmas ************************************************)
+
+lemma sor_inv_ppn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f2,f. ↑f1 = g1 → ↑f2 = g2 → ⫯f = g → ⊥.
+#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0
+elim (sor_inv_ppx … H … H1 H2) -g1 -g2 #x #_ #H destruct
+/2 width=3 by discr_push_next/
+qed-.
+
+lemma sor_inv_nxp: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f. ⫯f1 = g1 → ↑f = g → ⊥.
+#g1 #g2 #g #H #f1 #f #H1 #H0
+elim (pn_split g2) * #f2 #H2
+[ elim (sor_inv_npx … H … H1 H2)
+| elim (sor_inv_nnx … H … H1 H2)
+] -g1 -g2 #x #_ #H destruct
+/2 width=3 by discr_next_push/
+qed-.
+
+lemma sor_inv_xnp: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f2,f. ⫯f2 = g2 → ↑f = g → ⊥.
+#g1 #g2 #g #H #f2 #f #H2 #H0
+elim (pn_split g1) * #f1 #H1
+[ elim (sor_inv_pnx … H … H1 H2)
+| elim (sor_inv_nnx … H … H1 H2)
+] -g1 -g2 #x #_ #H destruct
+/2 width=3 by discr_next_push/
+qed-.
+
+lemma sor_inv_ppp: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f2,f. ↑f1 = g1 → ↑f2 = g2 → ↑f = g → f1 ⋓ f2 ≡ f.
+#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0
+elim (sor_inv_ppx … H … H1 H2) -g1 -g2 #x #Hx #H destruct
+<(injective_push … H) -f //
+qed-.
+
+lemma sor_inv_npn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f2,f. ⫯f1 = g1 → ↑f2 = g2 → ⫯f = g → f1 ⋓ f2 ≡ f.
+#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0
+elim (sor_inv_npx … H … H1 H2) -g1 -g2 #x #Hx #H destruct
+<(injective_next … H) -f //
+qed-.
+
+lemma sor_inv_pnn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f2,f. ↑f1 = g1 → ⫯f2 = g2 → ⫯f = g → f1 ⋓ f2 ≡ f.
+#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0
+elim (sor_inv_pnx … H … H1 H2) -g1 -g2 #x #Hx #H destruct
+<(injective_next … H) -f //
+qed-.
+
+lemma sor_inv_nnn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f2,f. ⫯f1 = g1 → ⫯f2 = g2 → ⫯f = g → f1 ⋓ f2 ≡ f.
+#g1 #g2 #g #H #f1 #f2 #f #H1 #H2 #H0
+elim (sor_inv_nnx … H … H1 H2) -g1 -g2 #x #Hx #H destruct
+<(injective_next … H) -f //
+qed-.
+
+lemma sor_inv_pxp: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f. ↑f1 = g1 → ↑f = g →
+                   ∃∃f2. f1 ⋓ f2 ≡ f & ↑f2 = g2.
+#g1 #g2 #g #H #f1 #f #H1 #H0
+elim (pn_split g2) * #f2 #H2
+[ /3 width=7 by sor_inv_ppp, ex2_intro/
+| elim (sor_inv_xnp … H … H2 H0)
+]
+qed-.
+
+lemma sor_inv_xpp: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f2,f. ↑f2 = g2 → ↑f = g →
+                   ∃∃f1. f1 ⋓ f2 ≡ f & ↑f1 = g1.
+#g1 #g2 #g #H #f2 #f #H2 #H0
+elim (pn_split g1) * #f1 #H1
+[ /3 width=7 by sor_inv_ppp, ex2_intro/
+| elim (sor_inv_nxp … H … H1 H0)
+]
+qed-.
+
+lemma sor_inv_pxn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f. ↑f1 = g1 → ⫯f = g →
+                   ∃∃f2. f1 ⋓ f2 ≡ f & ⫯f2 = g2.
+#g1 #g2 #g #H #f1 #f #H1 #H0
+elim (pn_split g2) * #f2 #H2
+[ elim (sor_inv_ppn … H … H1 H2 H0)
+| /3 width=7 by sor_inv_pnn, ex2_intro/
+]
+qed-.
+
+lemma sor_inv_xpn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f2,f. ↑f2 = g2 → ⫯f = g →
+                   ∃∃f1. f1 ⋓ f2 ≡ f & ⫯f1 = g1.
+#g1 #g2 #g #H #f2 #f #H2 #H0
+elim (pn_split g1) * #f1 #H1
+[ elim (sor_inv_ppn … H … H1 H2 H0)
+| /3 width=7 by sor_inv_npn, ex2_intro/
+]
+qed-.
+
+lemma sor_inv_xxp: ∀g1,g2,g. g1 ⋓ g2 ≡ g → ∀f. ↑f = g →
+                   ∃∃f1,f2. f1 ⋓ f2 ≡ f & ↑f1 = g1 & ↑f2 = g2.
+#g1 #g2 #g #H #f #H0
+elim (pn_split g1) * #f1 #H1
+[ elim (sor_inv_pxp … H … H1 H0) -g /2 width=5 by ex3_2_intro/
+| elim (sor_inv_nxp … H … H1 H0)
+]
+qed-.
+
+lemma sor_inv_nxn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f1,f. ⫯f1 = g1 → ⫯f = g →
+                   (∃∃f2. f1 ⋓ f2 ≡ f & ↑f2 = g2) ∨
+                   ∃∃f2. f1 ⋓ f2 ≡ f & ⫯f2 = g2.
+#g1 #g2 elim (pn_split g2) *
+/4 width=7 by sor_inv_npn, sor_inv_nnn, ex2_intro, or_intror, or_introl/
+qed-.
+
+lemma sor_inv_xnn: ∀g1,g2,g. g1 ⋓ g2 ≡ g →
+                   ∀f2,f. ⫯f2 = g2 → ⫯f = g →
+                   (∃∃f1. f1 ⋓ f2 ≡ f & ↑f1 = g1) ∨
+                   ∃∃f1. f1 ⋓ f2 ≡ f & ⫯f1 = g1.
+#g1 elim (pn_split g1) *
+/4 width=7 by sor_inv_pnn, sor_inv_nnn, ex2_intro, or_intror, or_introl/
+qed-.
+
+lemma sor_inv_xxn: ∀g1,g2,g. g1 ⋓ g2 ≡ g → ∀f. ⫯f = g →
+                   ∨∨ ∃∃f1,f2. f1 ⋓ f2 ≡ f & ⫯f1 = g1 & ↑f2 = g2
+                    | ∃∃f1,f2. f1 ⋓ f2 ≡ f & ↑f1 = g1 & ⫯f2 = g2
+                    | ∃∃f1,f2. f1 ⋓ f2 ≡ f & ⫯f1 = g1 & ⫯f2 = g2.
+#g1 #g2 #g #H #f #H0
+elim (pn_split g1) * #f1 #H1
+[ elim (sor_inv_pxn … H … H1 H0) -g
+  /3 width=5 by or3_intro1, ex3_2_intro/
+| elim (sor_inv_nxn … H … H1 H0) -g *
+  /3 width=5 by or3_intro0, or3_intro2, ex3_2_intro/
+]
+qed-.
+
 (* Main inversion lemmas ****************************************************)
 
 corec theorem sor_mono: ∀f1,f2,x,y. f1 ⋓ f2 ≡ x → f1 ⋓ f2 ≡ y → x ≗ y.
@@ -131,7 +266,27 @@ 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 identity ****************************************************)
+(* 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.
+
+(* 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
@@ -145,17 +300,47 @@ corec lemma sor_isid_dx: ∀f2. 𝐈⦃f2⦄ → ∀f1. f1 ⋓ f2 ≡ f1.
 /3 width=7 by sor_pp, sor_np/
 qed.
 
-(* Properties on finite colength assignment *********************************)
+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 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/
+qed-.
+
+lemma sor_isid_inv_dx: ∀f1,f2,f. f1 ⋓ f2 ≡ f → 𝐈⦃f2⦄ → f1 ≗ f.
+/3 width=4 by sor_isid_dx, sor_mono/
+qed-.
+
+corec lemma sor_fwd_isid1: ∀f1,f2,f. f1 ⋓ f2 ≡ f → 𝐈⦃f⦄ → 𝐈⦃f1⦄.
+#f1 #f2 #f * -f1 -f2 -f
+#f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H #Hg
+[ /4 width=6 by isid_inv_push, isid_push/ ]
+cases (isid_inv_next … Hg … H)
+qed-.
+
+corec lemma sor_fwd_isid2: ∀f1,f2,f. f1 ⋓ f2 ≡ f → 𝐈⦃f⦄ → 𝐈⦃f2⦄.
+#f1 #f2 #f * -f1 -f2 -f
+#f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H #Hg
+[ /4 width=6 by isid_inv_push, isid_push/ ]
+cases (isid_inv_next … Hg … H)
+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 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.
 #f1 #n1 #Hf1 elim Hf1 -f1 -n1 /3 width=6 by sor_isid_sn, ex4_2_intro/
 #f1 #n1 #Hf1 #IH #f2 #n2 * -f2 -n2 /3 width=6 by fcla_push, fcla_next, ex4_2_intro, sor_isid_dx/
-#f2 #n2 #Hf2 elim (IH … Hf2) -IH -Hf2 -Hf1
-[ /3 width=7 by fcla_push, sor_pp, ex4_2_intro/
-| /3 width=7 by fcla_next, sor_pn, max_S2_le_S, le_S_S, ex4_2_intro/
-| /3 width=7 by fcla_next, sor_np, max_S1_le_S, le_S_S, ex4_2_intro/
+#f2 #n2 #Hf2 elim (IH … Hf2) -IH -Hf2 -Hf1 [2,4: #f #n <plus_n_Sm ] (**) (* full auto fails *)
+[ /3 width=7 by fcla_next, sor_pn, max_S2_le_S, le_S_S, ex4_2_intro/
 | /4 width=7 by fcla_next, sor_nn, le_S, le_S_S, ex4_2_intro/
+| /3 width=7 by fcla_push, sor_pp, ex4_2_intro/
+| /3 width=7 by fcla_next, sor_np, max_S1_le_S, le_S_S, ex4_2_intro/
 ]
 qed-.
 
@@ -165,7 +350,26 @@ lemma sor_fcla: ∀f1,n1. 𝐂⦃f1⦄ ≡ n1 → ∀f2,n2. 𝐂⦃f2⦄ ≡ n2
 /4 width=6 by sor_mono, fcla_eq_repl_back, ex3_intro/
 qed-.
 
-(* Properties on test for finite colength ***********************************)
+(* Forward lemmas with finite colength **************************************)
+
+lemma sor_fwd_fcla_sn_ex: ∀f,n. 𝐂⦃f⦄ ≡ n → ∀f1,f2. f1 ⋓ f2 ≡ f →
+                          ∃∃n1.  𝐂⦃f1⦄ ≡ n1 & n1 ≤ n.
+#f #n #H elim H -f -n
+[ /4 width=4 by sor_fwd_isid1, fcla_isid, ex2_intro/
+| #f #n #_ #IH #f1 #f2 #H
+  elim (sor_inv_xxp … H) -H [ |*: // ] #g1 #g2 #Hf #H1 #H2 destruct
+  elim (IH … Hf) -f /3 width=3 by fcla_push, ex2_intro/
+| #f #n #_ #IH #f1 #f2 #H
+  elim (sor_inv_xxn … H) -H [1,3,4: * |*: // ] #g1 #g2 #Hf #H1 #H2 destruct
+  elim (IH … Hf) -f /3 width=3 by fcla_push, fcla_next, le_S_S, le_S, ex2_intro/
+]
+qed-.
+
+lemma sor_fwd_fcla_dx_ex: ∀f,n. 𝐂⦃f⦄ ≡ n → ∀f1,f2. f1 ⋓ f2 ≡ f →
+                          ∃∃n2.  𝐂⦃f2⦄ ≡ n2 & n2 ≤ n.
+/3 width=4 by sor_fwd_fcla_sn_ex, sor_sym/ qed-.
+
+(* 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
@@ -176,3 +380,42 @@ lemma sor_isfin: ∀f1,f2. 𝐅⦃f1⦄ → 𝐅⦃f2⦄ → ∀f. f1 ⋓ f2 ≡
 #f1 #f2 #Hf1 #Hf2 #f #Hf elim (sor_isfin_ex … Hf1 … Hf2) -Hf1 -Hf2
 /3 width=6 by sor_mono, isfin_eq_repl_back/
 qed-.
+
+(* Forward lemmas with test for finite colength *****************************)
+
+lemma sor_fwd_isfin_sn: ∀f. 𝐅⦃f⦄ → ∀f1,f2. f1 ⋓ f2 ≡ f → 𝐅⦃f1⦄.
+#f * #n #Hf #f1 #f2 #H
+elim (sor_fwd_fcla_sn_ex … Hf … H) -f -f2 /2 width=2 by ex_intro/
+qed-.
+
+lemma sor_fwd_isfin_dx: ∀f. 𝐅⦃f⦄ → ∀f1,f2. f1 ⋓ f2 ≡ f → 𝐅⦃f2⦄.
+#f * #n #Hf #f1 #f2 #H
+elim (sor_fwd_fcla_dx_ex … Hf … H) -f -f1 /2 width=2 by ex_intro/
+qed-.
+
+(* Inversion lemmas with test for finite colength ***************************)
+
+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 with inclusion ******************************************)
+
+corec lemma sor_inv_sle_sn: ∀f1,f2,f. f1 ⋓ f2 ≡ f → f1 ⊆ f.
+#f1 #f2 #f * -f1 -f2 -f
+#f1 #f2 #f #g1 #g2 #g #Hf #H1 #H2 #H0
+/3 width=5 by sle_push, sle_next, sle_weak/
+qed-.
+
+corec lemma sor_inv_sle_dx: ∀f1,f2,f. f1 ⋓ f2 ≡ f → f2 ⊆ f.
+#f1 #f2 #f * -f1 -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.