]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma
refactoring ...
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / arith.ma
index d181e88cf3f385f1e8ae993f172461810ca31036..9a3f0810f14ae3e512eebae395e07c8ea89ff565 100644 (file)
@@ -15,7 +15,7 @@
 include "ground_2/notation/functions/successor_1.ma".
 include "ground_2/notation/functions/predecessor_1.ma".
 include "arithmetics/nat.ma".
-include "ground_2/lib/star.ma".
+include "ground_2/lib/relations.ma".
 
 (* ARITHMETICAL PROPERTIES **************************************************)
 
@@ -23,6 +23,10 @@ interpretation "nat successor" 'Successor m = (S m).
 
 interpretation "nat predecessor" 'Predecessor m = (pred m).
 
+interpretation "nat min" 'and x y = (min x y).
+
+interpretation "nat max" 'or x y = (max x y).
+
 (* Iota equations ***********************************************************)
 
 lemma pred_O: pred 0 = 0.
@@ -31,8 +35,25 @@ normalize // qed.
 lemma pred_S: ∀m. pred (S m) = m.
 // qed.
 
+lemma plus_S1: ∀x,y. ⫯(x+y) = (⫯x) + y.
+// qed.
+
+lemma max_O1: ∀n. n = (0 ∨ n).
+// qed.
+
+lemma max_O2: ∀n. n = (n ∨ 0).
+// qed.
+
+lemma max_SS: ∀n1,n2. ⫯(n1∨n2) = (⫯n1 ∨ ⫯n2).
+#n1 #n2 elim (decidable_le n1 n2) #H normalize
+[ >(le_to_leb_true … H) | >(not_le_to_leb_false … H) ] -H //
+qed.
+
 (* Equations ****************************************************************)
 
+lemma plus_SO: ∀n. n + 1 = ⫯n.
+// qed.
+
 lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m.
 // qed-.
 
@@ -75,6 +96,20 @@ 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 idempotent_max: ∀n:nat. n = (n ∨ n).
+#n normalize >le_to_leb_true //
+qed.
+
+lemma associative_max: associative … max.
+#x #y #z normalize
+@(leb_elim x y) normalize #Hxy
+@(leb_elim y z) normalize #Hyz //
+[1,2: >le_to_leb_true /2 width=3 by transitive_le/
+| >not_le_to_leb_false /4 width=3 by lt_to_not_le, not_le_to_lt, transitive_lt/
+  >not_le_to_leb_false //
+]
+qed.
+
 (* Properties ***************************************************************)
 
 lemma eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2).
@@ -109,6 +144,15 @@ qed.
 lemma lt_S_S: ∀x,y. x < y → ⫯x < ⫯y.
 /2 width=1 by le_S_S/ qed.
 
+lemma lt_S: ∀n,m. n < m → n < ⫯m.
+/2 width=1 by le_S/ qed.
+
+lemma max_S1_le_S: ∀n1,n2,n. (n1 ∨ n2) ≤ n → (⫯n1 ∨ n2) ≤ ⫯n.
+/4 width=2 by to_max, le_maxr, le_S_S, le_S/ qed-.
+
+lemma max_S2_le_S: ∀n1,n2,n. (n1 ∨ n2) ≤ n → (n1 ∨ ⫯n2) ≤ ⫯n.
+/2 width=1 by max_S1_le_S/ qed-.
+
 lemma arith_j: ∀x,y,z. x-y-1 ≤ x-(y-z)-1.
 /3 width=1 by monotonic_le_minus_l, monotonic_le_minus_r/ qed.
 
@@ -128,9 +172,23 @@ 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-.
 
@@ -145,6 +203,25 @@ qed-.
 lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥.
 /3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-.
 
+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-.
@@ -190,10 +267,19 @@ lemma lt_elim: ∀R:relation nat.
 ]
 qed-.
 
+lemma le_elim: ∀R:relation nat.
+               (∀n2. R O (n2)) →
+               (∀n1,n2. R n1 n2 → R (⫯n1) (⫯n2)) →
+               ∀n1,n2. n1 ≤ n2 → R n1 n2.
+#R #IH1 #IH2 #n1 #n2 @(nat_elim2 … n1 n2) -n1 -n2
+/4 width=1 by monotonic_pred/ -IH1 -IH2
+#n1 #H elim (lt_le_false … H) -H //
+qed-.
+
 (* Iterators ****************************************************************)
 
 (* Note: see also: lib/arithemetics/bigops.ma *)
-let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝
+rec definition iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝
   match n with
    [ O   ⇒ nil
    | S k ⇒ op (iter k B op nil)
@@ -222,7 +308,7 @@ qed.
 (* Trichotomy operator ******************************************************)
 
 (* Note: this is "if eqb n1 n2 then a2 else if leb n1 n2 then a1 else a3" *)
-let rec tri (A:Type[0]) n1 n2 a1 a2 a3 on n1 : A ≝
+rec definition tri (A:Type[0]) n1 n2 a1 a2 a3 on n1 : A ≝
   match n1 with
   [ O    ⇒ match n2 with [ O ⇒ a2 | S n2 ⇒ a1 ]
   | S n1 ⇒ match n2 with [ O ⇒ a3 | S n2 ⇒ tri A n1 n2 a1 a2 a3 ]