]> matita.cs.unibo.it Git - helm.git/commitdiff
- work in progress proceeds for the new definition of voids ...
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 4 Dec 2017 20:19:29 +0000 (20:19 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 4 Dec 2017 20:19:29 +0000 (20:19 +0000)
- arith.ma: some additions and reordering

matita/matita/contribs/lambdadelta/basic_2/static/fle.ma
matita/matita/contribs/lambdadelta/basic_2/static/fle_fqup.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/voids.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/voids_length.ma
matita/matita/contribs/lambdadelta/basic_2/syntax/voids_voids.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl
matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma

index 7340a8abc5b8ee924dc7a1809128481f13f46003..02d3ecdb118d27724e6bbd33c2a3f7be4ede17d5 100644 (file)
 
 include "ground_2/relocation/rtmap_id.ma".
 include "basic_2/notation/relations/subseteq_4.ma".
-include "basic_2/syntax/voids_length.ma".
+include "basic_2/syntax/voids.ma".
 include "basic_2/static/frees.ma".
 
 (* FREE VARIABLES INCLUSION FOR RESTRICTED CLOSURES *************************)
 
-inductive fle (T1) (T2): relation lenv ≝
-| fle_intro: ∀f1,f2,L1,L2,n1,n2. ⓧ*[n1]L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 → ⓧ*[n2]L2 ⊢ 𝐅*⦃T2⦄ ≡ f2 →
-             |L1| = |L2| → ⫱*[n1]f1 ⊆ ⫱*[n2]f2 → fle T1 T2 (ⓧ*[n1]L1) (ⓧ*[n2]L2)
-.
+definition fle: bi_relation lenv term ≝ λL1,T1,L2,T2.
+                ∃∃n1,n2,f1,f2. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 & L2 ⊢ 𝐅*⦃T2⦄ ≡ f2 &
+                               ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 & ⫱*[n1]f1 ⊆ ⫱*[n2]f2.
 
 interpretation "free variables inclusion (restricted closure)"
-   'SubSetEq L1 T1 L2 T2 = (fle T1 T2 L1 L2).
+   'SubSetEq L1 T1 L2 T2 = (fle L1 T1 L2 T2).
 
 (* Basic properties *********************************************************)
 
-lemma fle_sort: ∀L1,L2. |L1| = |L2| → ∀s1,s2. ⦃L1, ⋆s1⦄ ⊆ ⦃L2, ⋆s2⦄.
-/3 width=5 by frees_sort, sle_refl, fle_intro/ qed.
+lemma fle_sort: ∀L,s1,s2. ⦃L, ⋆s1⦄ ⊆ ⦃L, ⋆s2⦄.
+#L elim (voids_refl L)
+/3 width=8 by frees_sort, sle_refl, ex4_4_intro/
+qed.
 
-lemma fle_gref: ∀L1,L2. |L1| = |L2| → ∀l1,l2. ⦃L1, §l1⦄ ⊆ ⦃L2, §l2⦄.
-/3 width=5 by frees_gref, sle_refl, fle_intro/ qed.
+lemma fle_gref: ∀L,l1,l2. ⦃L, §l1⦄ ⊆ ⦃L, §l2⦄.
+#L elim (voids_refl L)
+/3 width=8 by frees_gref, sle_refl, ex4_4_intro/
+qed.
 
 (* Basic inversion lemmas ***************************************************)
-
+(*
 fact fle_inv_voids_aux: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄ →
                         ∀K1,K2,n1,n2. |K1| = |K2| → L1 = ⓧ*[n1]K1 → L2 = ⓧ*[n2]K2 →
                         ∃∃f1,f2. ⓧ*[n1]K1 ⊢ 𝐅*⦃T1⦄ ≡ f1 & ⓧ*[n2]K2 ⊢ 𝐅*⦃T2⦄ ≡ f2 & ⫱*[n1]f1 ⊆ ⫱*[n2]f2.
@@ -52,3 +55,4 @@ qed-.
 lemma fle_inv_voids_sn: ∀L1,L2,T1,T2,n. ⦃ⓧ*[n]L1, T1⦄ ⊆ ⦃L2, T2⦄ → |L1| = |L2| →
                         ∃∃f1,f2. ⓧ*[n]L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 & L2 ⊢ 𝐅*⦃T2⦄ ≡ f2 & ⫱*[n]f1 ⊆ f2.
 /2 width=3 by fle_inv_voids_sn_aux/ qed-.
+*)
\ No newline at end of file
index aeed5c14632b0e9eed36f5a226df5827e1ff3b82..6ecf55f0a83a0d703783e6dcc3aeaa36c5e28a73 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+include "basic_2/syntax/voids_length.ma".
 include "basic_2/static/frees_fqup.ma".
 include "basic_2/static/fle.ma".
 
 (* FREE VARIABLES INCLUSION FOR RESTRICTED CLOSURES *************************)
 
 (* Advanced properties ******************************************************)
-(*
+
 lemma fle_refl: bi_reflexive … fle.
-#L #T elim (frees_total L T) /2 width=5 by sle_refl, ex3_2_intro/
+#L #T
+elim (voids_refl L) #n #Hn
+elim (frees_total L T) #f #Hf
+/2 width=8 by sle_refl, ex4_4_intro/
 qed.
-*)
+
 lemma fle_bind_dx_sn: ∀L1,L2,V1,V2. ⦃L1, V1⦄ ⊆ ⦃L2, V2⦄ →
                       ∀p,I,T2. ⦃L1, V1⦄ ⊆ ⦃L2, ⓑ{p,I}V2.T2⦄.
-#L1 #L2 #V1 #V2 * -L1 #f1 #g1 #L1 #n #Hf1 #Hg1 #HL12 #Hfg1 #p #I #T2
+#L1 #L2 #V1 #V2 * #n1 #m1 #f1 #g1 #Hf1 #Hg1 #HL12 #Hfg1 #p #I #T2
 elim (frees_total (L2.ⓧ) T2) #g2 #Hg2
 elim (sor_isfin_ex g1 (⫱g2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #g #Hg #_
-/4 width=8 by fle_intro, frees_bind_void, sor_inv_sle_sn, sle_trans/
+@(ex4_4_intro … g Hf1 … HL12) (**) (* full auto too slow *)
+/4 width=5 by frees_bind_void, sor_inv_sle_sn, sor_tls, sle_trans/
 qed.
-(*
-lemma fle_bind_dx_dx: ∀L1,L2,T1,T2. ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2.ⓧ, T2⦄ →
-                      ∀p,I,V2. ⦃L1.ⓧ, T1⦄ ⊆ ⦃L2, ⓑ{p,I}V2.T2⦄.
-#L1 #L2 #T1 #T2 * -L1 #f2 #g2 #L1 #n #Hf2 #Hg2 #HL12 #Hfg2 #p #I #V2
+
+lemma fle_bind_dx_dx: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L2.ⓧ, T2⦄ → |L1| ≤ |L2| →
+                      ∀p,I,V2. ⦃L1, T1⦄ ⊆ ⦃L2, ⓑ{p,I}V2.T2⦄.
+#L1 #L2 #T1 #T2 * #n1 #x1 #f2 #g2 #Hf2 #Hg2 #H #Hfg2 #HL12 #p #I #V2
+elim (voids_inv_void_dx_length … H HL12) -H -HL12 #m1 #HL12 #H #_ destruct
+<tls_xn in Hfg2; #Hfg2
 elim (frees_total L2 V2) #g1 #Hg1
 elim (sor_isfin_ex g1 (⫱g2)) /3 width=3 by frees_fwd_isfin, isfin_tl/ #g #Hg #_
-@(fle_intro … g … Hf2) /2 width=5 by frees_bind_void/
-@(sle_trans … Hfg1) @(sor_inv_sle_sn … Hg)
-
-
-
-/4 width=8 by fle_intro, frees_bind_void, sor_inv_sle_dx, sle_trans/
+@(ex4_4_intro … g Hf2 … HL12) (**) (* full auto too slow *)
+/4 width=5 by frees_bind_void, sor_inv_sle_dx, sor_tls, sle_trans/
 qed.
-*)
+
 lemma fle_flat_dx_sn: ∀L1,L2,V1,V2. ⦃L1, V1⦄ ⊆ ⦃L2, V2⦄ →
                       ∀I,T2. ⦃L1, V1⦄ ⊆ ⦃L2, ⓕ{I}V2.T2⦄.
-#L1 #L2 #V1 #V2 * -L1 #f1 #g1 #L1 #n #Hf1 #Hg1 #HL12 #Hfg1 #I #T2
+#L1 #L2 #V1 #V2 * #n1 #m1 #f1 #g1 #Hf1 #Hg1 #HL12 #Hfg1 #I #T2
 elim (frees_total L2 T2) #g2 #Hg2
 elim (sor_isfin_ex g1 g2) /2 width=3 by frees_fwd_isfin/ #g #Hg #_
-/4 width=8 by fle_intro, frees_flat, sor_inv_sle_sn, sle_trans/
+@(ex4_4_intro … g Hf1 … HL12) (**) (* full auto too slow *)
+/4 width=5 by frees_flat, sor_inv_sle_sn, sor_tls, sle_trans/
 qed.
 
 lemma fle_flat_dx_dx: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊆ ⦃L2, T2⦄ →
                       ∀I,V2. ⦃L1, T1⦄ ⊆ ⦃L2, ⓕ{I}V2.T2⦄.
-#L1 #L2 #T1 #T2 * -L1 #f2 #g2 #L1 #n #Hf2 #Hg2 #HL12 #Hfg2 #I #V2
+#L1 #L2 #T1 #T2 * #n1 #m1 #f2 #g2 #Hf2 #Hg2 #HL12 #Hfg2 #I #V2
 elim (frees_total L2 V2) #g1 #Hg1
 elim (sor_isfin_ex g1 g2) /2 width=3 by frees_fwd_isfin/ #g #Hg #_
-/4 width=8 by fle_intro, frees_flat, sor_inv_sle_dx, sle_trans/
+@(ex4_4_intro … g Hf2 … HL12) (**) (* full auto too slow *)
+/4 width=5 by frees_flat, sor_inv_sle_dx, sor_tls, sle_trans/
 qed.
index 8c58fd478cb203bfebee9ad77a523b30994162b6..ca5c4fc58a691c8537e4c7d55a9780147e41743a 100644 (file)
@@ -45,30 +45,19 @@ lemma voids_sym: bi_symmetric … voids.
 /2 width=2 by voids_atom, voids_pair_sn, voids_pair_dx, voids_void_sn, voids_void_dx/
 qed-.
 
-(*
-
-
-
-
-lemma voids_zero: ∀L. L = ⓧ*[0]L.
-// qed.
-
-lemma voids_succ: ∀L,n. (ⓧ*[n]L).ⓧ = ⓧ*[⫯n]L.
-// qed.
-
-(* Advanced properties ******************************************************)
-
-lemma voids_next: ∀n,L. ⓧ*[n](L.ⓧ) = ⓧ*[⫯n]L.
-#n elim n -n //
-qed.
-
-(* Main inversion properties ************************************************)
-
-theorem voids_inj: ∀n. injective … (λL. ⓧ*[n]L).
-#n elim n -n //
-#n #IH #L1 #L2
-<voids_succ <voids_succ #H
-elim (destruct_lbind_lbind_aux … H) -H (**) (* destruct lemma needed *)
-/2 width=1 by/
+(* Advanced Inversion lemmas ************************************************)
+
+fact voids_inv_void_dx_aux: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 →
+                            ∀K2,m2. n2 = ⫯m2 → L2 = K2.ⓧ → ⓧ*[n1]L1 ≋ ⓧ*[m2]K2.
+#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2
+[ #K2 #m2 #H destruct
+| #I1 #I2 #L1 #L2 #V #n #_ #_ #K2 #m2 #H destruct
+| #I1 #I2 #L1 #L2 #V #n #_ #_ #K2 #m2 #H destruct
+| #L1 #L2 #n1 #n2 #_ #IH #K2 #m2 #H1 #H2 destruct
+  /3 width=1 by voids_void_sn/
+| #L1 #L2 #n1 #n2 #HL12 #_ #K2 #m2 #H1 #H2 destruct //
+]
 qed-.
-*)
\ No newline at end of file
+
+lemma voids_inv_void_dx: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[⫯n2]L2.ⓧ → ⓧ*[n1]L1 ≋ ⓧ*[n2]L2.
+/2 width=5 by voids_inv_void_dx_aux/ qed-.
index 30738ee24709f3b42e7ddb6d41bf41de4c0b1e42..60d41924ba064ea6e45437c45da4cc49f4d42e84 100644 (file)
@@ -19,25 +19,40 @@ include "basic_2/syntax/voids.ma".
 
 (* Forward lemmas with length for local environments ************************)
 
+lemma voids_fwd_length_le_sn: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 → n1 ≤ |L1|.
+#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 normalize
+/2 width=1 by le_S_S/
+qed-.
+
+lemma voids_fwd_length_le_dx: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 → n2 ≤ |L2|.
+#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 normalize
+/2 width=1 by le_S_S/
+qed-.
+
 lemma voids_fwd_length: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 →
                         |L1| + n2 = |L2| + n1.
-#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 normalize //
-#I1 #I2 #K1 #K2 #V #n #_ #IH 
-
-(* Main forward properties with length for local environments ***************)
-
-theorem voids_inj_length: ∀n1,n2,L1,L2. ⓧ*[n1]L1 = ⓧ*[n2]L2 →
-                          |L1| = |L2| → n1 = n2 ∧ L1 = L2.
-#n1 elim n1 -n1
-[ * /2 width=1 by conj/ #n2 #L1 #L2 | #n1 #IH * [ | #n2 ] #L1 #L2 ]
-[ <voids_zero #H destruct
-  <length_void <commutative_plus #H
-  elim (plus_xSy_x_false … H)
-| <voids_zero #H destruct
-  <length_void <commutative_plus #H
-  elim (plus_xSy_x_false … (sym_eq … H))
-| <voids_succ <voids_succ #H #HL12
-  elim (destruct_lbind_lbind_aux … H) -H (**) (* destruct lemma needed *)
-  #H #_ elim (IH … H HL12) -IH -H -HL12 /2 width=1 by conj/
-]
+#L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2 normalize
+/2 width=2 by injective_plus_r/
+qed-.
+
+lemma voids_fwd_length_minus: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 →
+                              |L1| - n1 = |L2| - n2.
+/3 width=3 by voids_fwd_length, voids_fwd_length_le_dx, voids_fwd_length_le_sn, plus_to_minus_2/ qed-.
+
+lemma voids_inj_length: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 →
+                        |L1| = |L2| → n1 = n2.
+#L1 #L2 #n1 #n2 #H #HL12
+lapply (voids_fwd_length … H) -H #H
+/2 width=2 by injective_plus_l/
+qed-.
+
+(* Inversion lemmas with length for local environments **********************)
+
+lemma voids_inv_void_dx_length: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2.ⓧ → |L1| ≤ |L2| →
+                                ∃∃m2. ⓧ*[n1]L1 ≋ ⓧ*[m2]L2 & n2 = ⫯m2 & n1 ≤ m2.
+#L1 #L2 #n1 #n2 #H #HL12
+lapply (voids_fwd_length … H) normalize >plus_n_Sm #H0
+lapply (plus2_inv_le_sn … H0 HL12) -H0 -HL12 #H0
+elim (le_inv_S1 … H0) -H0 #m2 #Hm2 #H0 destruct
+/3 width=4 by voids_inv_void_dx, ex3_intro/
 qed-.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/syntax/voids_voids.ma b/matita/matita/contribs/lambdadelta/basic_2/syntax/voids_voids.ma
new file mode 100644 (file)
index 0000000..5ee9d47
--- /dev/null
@@ -0,0 +1,23 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "basic_2/syntax/voids.ma".
+
+(* EQUIVALENCE FOR LOCAL ENVIRONMENTS UP TO EXCLUSION BINDERS ***************)
+
+(* Main inversion lemmas ****************************************************)
+
+theorem voids_inj: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ ⓧ*[n2]L2 →
+                   ∀m1,m2. ⓧ*[m1]L1 ≋ ⓧ*[m2]L2 →
+                   ∧∧ n1 = m1 & n2 = m2.
index 4368142cb476779fc4b97d676daa16d1559d7eb5..0dd04cb2cc7f415862dd4cf324671f220b7eb7ac 100644 (file)
@@ -238,6 +238,11 @@ table {
    ]
    class "red"
    [ { "syntax" * } {
+        [ { "equivalence up to exclusion binders" * } {
+             [ [ "" ] "voids ( ⓧ*[?]? ≋ ⓧ*[?]? )" "voids_length" + "voids_voids" * ]
+          }
+        ]
+   
         [ { "append for local environments" * } {
              [ [ "" ] "append ( ? @@ ? )" "append_length" * ]
           }
@@ -263,7 +268,6 @@ table {
         [ { "local environments" * } {
              [ [ "" ] "ceq_ext" "ceq_ext_ceq_ext" * ]
              [ [ "" ] "cext2" * ]
-             [ [ "" ] "lenv_voids ( ⓧ*[?]? )" * ]
              [ [ "" ] "lenv_length ( |?| )" * ]
              [ [ "" ] "lenv_weight ( ♯{?} )" * ]
              [ [ "" ] "lenv" * ]
index d9355a1a04cba417f48c68552df06dceed8f1a2e..e3e22484394c91625310912317448ae62121db81 100644 (file)
@@ -69,21 +69,21 @@ lemma plus_n_2: ∀n. n + 2 = n + 1 + 1.
 // qed.
 
 lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p).
-/2 by plus_minus/ qed.
+/2 by plus_minus/ qed-.
 
 lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → m + n - p = m - p + n.
-/2 by plus_minus/ qed.
+/2 by plus_minus/ qed-.
 
 lemma minus_minus_comm3: ∀n,x,y,z. n-x-y-z = n-y-z-x.
 // qed.
 
 lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b.
 #a #b #c1 #H >minus_minus_comm >minus_le_minus_minus_comm //
-qed.
+qed-.
 
 lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b.
 #a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1 by arith_b1/
-qed.
+qed-.
 
 lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b.
 /3 by monotonic_le_minus_l, le_to_le_to_eq, le_n/ qed.
@@ -91,11 +91,15 @@ lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b.
 lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b →
                 a1 - c1 + a2 - (b - c1) = a1 + a2 - b.
 #a1 #a2 #b #c1 #H1 #H2 >plus_minus /2 width=1 by arith_b2/
-qed.
+qed-.
 
 lemma arith_i: ∀x,y,z. y < x → x+z-y-1 = x-y-1+z.
 /2 width=1 by plus_minus/ qed-.
 
+lemma plus_to_minus_2: ∀m1,m2,n1,n2. n1 ≤ m1 → n2 ≤ m2 →
+                       m1+n2 = m2+n1 → m1-n1 = m2-n2.
+/2 width=1 by arith_b1/ qed-.
+
 lemma idempotent_max: ∀n:nat. n = (n ∨ n).
 #n normalize >le_to_leb_true //
 qed.
@@ -136,7 +140,7 @@ lemma monotonic_le_minus_l2: ∀x1,x2,y,z. x1 ≤ x2 → x1 - y - z ≤ x2 - y -
 /3 width=1 by monotonic_le_minus_l/ qed.
 
 (* Note: this might interfere with nat.ma *)
-lemma monotonic_lt_pred: ∀m,n. m < n → O < m → pred m < pred n.
+lemma monotonic_lt_pred: ∀m,n. m < n → 0 < m → pred m < pred n.
 #m #n #Hmn #Hm whd >(S_pred … Hm)
 @le_S_S_to_le >S_pred /2 width=3 by transitive_lt/
 qed.
@@ -172,26 +176,6 @@ qed.
 
 (* Inversion & forward lemmas ***********************************************)
 
-lemma nat_split: ∀x. x = 0 ∨ ∃y. ⫯y = x.
-* /3 width=2 by ex_intro, or_introl, or_intror/
-qed-.
-
-lemma max_inv_O3: ∀x,y. (x ∨ y) = 0 → 0 = x ∧ 0 = y.
-/4 width=2 by le_maxr, le_maxl, le_n_O_to_eq, conj/
-qed-.
-
-lemma plus_inv_O3: ∀x,y. x + y = 0 → x = 0 ∧ y = 0.
-/2 width=1 by plus_le_0/ qed-.
-
-lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0.
-// qed-.
-
-lemma discr_plus_x_xy: ∀x,y. x = x + y → y = 0.
-/2 width=2 by le_plus_minus_comm/ qed-.
-
-lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y.
-/2 width=1 by monotonic_pred/ qed-.
-
 lemma lt_refl_false: ∀n. n < n → ⊥.
 #n #H elim (lt_to_not_eq … H) -H /2 width=1 by/
 qed-.
@@ -207,25 +191,6 @@ lemma succ_inv_refl_sn: ∀x. ⫯x = x → ⊥.
 #x #H @(lt_le_false x (⫯x)) //
 qed-.
 
-lemma lt_inv_O1: ∀n. 0 < n → ∃m. ⫯m = n.
-* /2 width=2 by ex_intro/
-#H cases (lt_le_false … H) -H //
-qed-.
-
-lemma lt_inv_S1: ∀m,n. ⫯m < n → ∃∃p. m < p & ⫯p = n.
-#m * /3 width=3 by lt_S_S_to_lt, ex2_intro/
-#H cases (lt_le_false … H) -H //
-qed-.
-
-lemma lt_inv_gen: ∀y,x. x < y → ∃∃z. x ≤ z & ⫯z = y.
-* /3 width=3 by le_S_S_to_le, ex2_intro/
-#x #H elim (lt_le_false … H) -H //
-qed-.
-
-lemma pred_inv_refl: ∀m. pred m = m → m = 0.
-* // normalize #m #H elim (lt_refl_false m) //
-qed-.
-
 lemma le_plus_xSy_O_false: ∀x,y. x + S y ≤ 0 → ⊥.
 #x #y #H lapply (le_n_O_to_eq … H) -H <plus_n_Sm #H destruct
 qed-.
@@ -241,6 +206,28 @@ lemma plus_xySz_x_false: ∀z,x,y. x + y + S z = x → ⊥.
 lemma plus_xSy_x_false: ∀y,x. x + S y = x → ⊥.
 /2 width=4 by plus_xySz_x_false/ qed-.
 
+lemma pred_inv_refl: ∀m. pred m = m → m = 0.
+* // normalize #m #H elim (lt_refl_false m) //
+qed-.
+
+lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0.
+// qed-.
+
+lemma discr_plus_x_xy: ∀x,y. x = x + y → y = 0.
+/2 width=2 by le_plus_minus_comm/ qed-.
+
+lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y.
+/2 width=1 by monotonic_pred/ qed-.
+
+lemma plus2_inv_le_sn: ∀m1,m2,n1,n2. m1 + n1 = m2 + n2 → m1 ≤ m2 → n2 ≤ n1.
+#m1 #m2 #n1 #n2 #H #Hm
+lapply (monotonic_le_plus_l n1 … Hm) -Hm >H -H
+/2 width=2 by le_plus_to_le/
+qed-.
+
+lemma lt_S_S_to_lt: ∀x,y. ⫯x < ⫯y → x < y.
+/2 width=1 by le_S_S_to_le/ qed-.
+
 (* Note this should go in nat.ma *)
 lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0.
 #x @(nat_ind_plus … x) -x /2 width=1 by or_introl/
@@ -250,12 +237,35 @@ lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0.
 #H destruct
 qed-.
 
+lemma lt_inv_O1: ∀n. 0 < n → ∃m. ⫯m = n.
+* /2 width=2 by ex_intro/
+#H cases (lt_le_false … H) -H //
+qed-.
+
+lemma lt_inv_S1: ∀m,n. ⫯m < n → ∃∃p. m < p & ⫯p = n.
+#m * /3 width=3 by lt_S_S_to_lt, ex2_intro/
+#H cases (lt_le_false … H) -H //
+qed-.
+
+lemma lt_inv_gen: ∀y,x. x < y → ∃∃z. x ≤ z & ⫯z = y.
+* /3 width=3 by le_S_S_to_le, ex2_intro/
+#x #H elim (lt_le_false … H) -H //
+qed-.
+
+lemma plus_inv_O3: ∀x,y. x + y = 0 → x = 0 ∧ y = 0.
+/2 width=1 by plus_le_0/ qed-.
+
+lemma max_inv_O3: ∀x,y. (x ∨ y) = 0 → 0 = x ∧ 0 = y.
+/4 width=2 by le_maxr, le_maxl, le_n_O_to_eq, conj/
+qed-.
+
 lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y.
 * /2 width=1 by conj/ #x #y normalize #H destruct
 qed-.
 
-lemma lt_S_S_to_lt: ∀x,y. ⫯x < ⫯y → x < y.
-/2 width=1 by le_S_S_to_le/ qed-.
+lemma nat_split: ∀x. x = 0 ∨ ∃y. ⫯y = x.
+* /3 width=2 by ex_intro, or_introl, or_intror/
+qed-.
 
 lemma lt_elim: ∀R:relation nat.
                (∀n2. R O (⫯n2)) →