(* Equalities ***************************************************************)
-lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m.
-// qed-.
-
-lemma plus_minus_m_m_commutative (n) (m): m ≤ n → n = m+(n-m).
-/2 width=1 by plus_minus_associative/ qed-.
-
-lemma plus_to_minus_2: ∀m1,m2,n1,n2. n1 ≤ m1 → n2 ≤ m2 →
- m1+n2 = m2+n1 → m1-n1 = m2-n2.
-#m1 #m2 #n1 #n2 #H1 #H2 #H
-@plus_to_minus >plus_minus_associative //
-qed-.
-
-(* Note: uses minus_minus_comm, minus_plus_m_m, commutative_plus, plus_minus *)
-lemma plus_minus_minus_be: ∀x,y,z. y ≤ z → z ≤ x → (x - z) + (z - y) = x - y.
-#x #z #y #Hzy #Hyx >plus_minus // >commutative_plus >plus_minus //
-qed-.
-
-lemma lt_succ_pred: ∀m,n. n < m → m = ↑↓m.
-#m #n #Hm >S_pred /2 width=2 by ltn_to_ltO/
-qed-.
+(*** plus_minus_plus_plus_l *) (**)
+lemma plus_minus_plus_plus_l: ∀z,x,y,h. z + (x + h) - (y + h) = z + x - y.
+#H1 #H2 #H3 #H4
+<nplus_assoc <nminus_plus_dx_bi // qed-.
fact plus_minus_minus_be_aux: ∀i,x,y,z. y ≤ z → z ≤ x → i = z - y → x - z + i = x - y.
/2 width=1 by plus_minus_minus_be/ qed-.
-lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p).
-/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-.
-
-lemma minus_minus_comm3: ∀n,x,y,z. n-x-y-z = n-y-z-x.
-// qed.
-
(* Properties ***************************************************************)
-lemma eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2).
-#n1 elim n1 -n1 [| #n1 #IHn1 ] * [2,4: #n2 ]
-[1,4: @or_intror #H destruct
-| elim (IHn1 n2) -IHn1 /3 width=1 by or_intror, or_introl/
-| /2 width=1 by or_introl/
-]
-qed-.
-
-lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m.
-#m #n elim (lt_or_ge m n) /2 width=1 by or3_intro0/
-#H elim H -m /2 width=1 by or3_intro1/
-#m #Hm * /3 width=1 by not_le_to_lt, le_S_S, or3_intro2/
-qed-.
+(*** lt_plus_Sn_r *) (**)
+lemma lt_plus_Sn_r: ∀a,x,n. a < a + x + ↑n.
+/2 width=1/ qed-.
lemma monotonic_le_minus_l2: ∀x1,x2,y,z. x1 ≤ x2 → x1 - y - z ≤ x2 - y - z.
/3 width=1 by monotonic_le_minus_l/ qed.
-lemma minus_le_trans_sn: ∀x1,x2. x1 ≤ x2 → ∀x. x1-x ≤ x2.
-/2 width=3 by transitive_le/ qed.
-
-lemma le_plus_to_minus_l: ∀a,b,c. a + b ≤ c → b ≤ c-a.
-/2 width=1 by le_plus_to_minus_r/
-qed-.
-
-lemma le_plus_to_minus_comm: ∀n,m,p. n ≤ p+m → n-p ≤ m.
-/2 width=1 by le_plus_to_minus/ qed-.
-
-lemma le_inv_S1: ∀m,n. ↑m ≤ n → ∃∃p. m ≤ p & ↑p = n.
-#m *
-[ #H lapply (le_n_O_to_eq … H) -H
- #H destruct
-| /3 width=3 by monotonic_pred, ex2_intro/
-]
-qed-.
-
-(* Note: this might interfere with nat.ma *)
-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.
-
-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 monotonic_lt_minus_r:
-∀p,q,n. q < n -> q < p → n-p < n-q.
-#p #q #n #Hn #H
-lapply (monotonic_le_minus_r … n H) -H #H
-@(le_to_lt_to_lt … H) -H
-/2 width=1 by lt_plus_to_minus/
-qed.
-
(* Inversion & forward lemmas ***********************************************)
-lemma lt_refl_false: ∀n. n < n → ⊥.
-#n #H elim (lt_to_not_eq … H) -H /2 width=1 by/
-qed-.
-
-lemma lt_zero_false: ∀n. n < 0 → ⊥.
-#n #H elim (lt_to_not_le … H) -H /2 width=1 by/
-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 le_dec (n) (m): Decidable (n≤m).
-#n elim n -n [ /2 width=1 by or_introl/ ]
-#n #IH * [ /3 width=2 by lt_zero_false, or_intror/ ]
-#m elim (IH m) -IH
-[ /3 width=1 by or_introl, le_S_S/
-| /4 width=1 by or_intror, le_S_S_to_le/
-]
-qed-.
-
-lemma succ_inv_refl_sn: ∀x. ↑x = x → ⊥.
-#x #H @(lt_le_false x (↑x)) //
-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-.
-
lemma le_plus_xySz_x_false: ∀y,z,x. x + y + S z ≤ x → ⊥.
#y #z #x elim x -x /3 width=1 by le_S_S_to_le/
#H elim (le_plus_xSy_O_false … H)
lemma plus_xySz_x_false: ∀z,x,y. x + y + S z = x → ⊥.
/2 width=4 by le_plus_xySz_x_false/ qed-.
-lemma plus_xSy_x_false: ∀y,x. x + S y = x → ⊥.
-/2 width=4 by plus_xySz_x_false/ qed-.
-
-lemma pred_inv_fix_sn: ∀x. ↓x = x → 0 = x.
-* // #x <pred_Sn #H
-elim (succ_inv_refl_sn x) //
-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 plus2_le_sn_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 plus2_le_sn_dx: ∀m1,m2,n1,n2. m1 + n1 = n2 + m2 → m1 ≤ m2 → n2 ≤ n1.
-/2 width=4 by plus2_le_sn_sn/ qed-.
-
-lemma plus2_le_dx_sn: ∀m1,m2,n1,n2. n1 + m1 = m2 + n2 → m1 ≤ m2 → n2 ≤ n1.
-/2 width=4 by plus2_le_sn_sn/ qed-.
-
-lemma plus2_le_dx_dx: ∀m1,m2,n1,n2. n1 + m1 = n2 + m2 → m1 ≤ m2 → n2 ≤ n1.
-/2 width=4 by plus2_le_sn_sn/ 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/
-#x #_ #y @(nat_ind_plus … y) -y /2 width=1 by or_intror/
-#y #_ >minus_plus_plus_l
-#H lapply (discr_plus_xy_minus_xz … H) -H
-#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 plus_inv_S3_sn: ∀x1,x2,x3. x1+x2 = ↑x3 →
- ∨∨ ∧∧ x1 = 0 & x2 = ↑x3
- | ∃∃y1. x1 = ↑y1 & y1 + x2 = x3.
-* /3 width=1 by or_introl, conj/
-#x1 #x2 #x3 <plus_S1 #H destruct
-/3 width=3 by ex2_intro, or_intror/
-qed-.
-
-lemma plus_inv_S3_dx: ∀x2,x1,x3. x1+x2 = ↑x3 →
- ∨∨ ∧∧ x2 = 0 & x1 = ↑x3
- | ∃∃y2. x2 = ↑y2 & x1 + y2 = x3.
-* /3 width=1 by or_introl, conj/
-#x2 #x1 #x3 <plus_n_Sm #H destruct
-/3 width=3 by ex2_intro, or_intror/
-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 nat_split: ∀x. x = 0 ∨ ∃y. ↑y = x.
-* /3 width=2 by ex_intro, or_introl, or_intror/
-qed-.
-
lemma nat_elim_le_sn (Q:relation …):
(∀m1,m2. (∀m. m < m2-m1 → Q (m2-m) m2) → m1 ≤ m2 → Q m1 m2) →
∀n1,n2. n1 ≤ n2 → Q n1 n2.
(* Basic constructions ******************************************************)
+(*** eq_nat_dec *)
lemma eq_nat_dec (n1,n2:nat): Decidable (n1 = n2).
* [| #p1 ] * [2,4: #p2 ]
[1,4: @or_intror #H destruct
-"cic:/matita/arithmetics/nat/nat.ind"
-#
-"cic:/matita/arithmetics/nat/nat_ind.con"
-"cic:/matita/arithmetics/nat/nat_elim2.con"
-"cic:/matita/arithmetics/nat/injective_S.con"
-#
-"cic:/matita/arithmetics/nat/plus.con"
-"cic:/matita/arithmetics/nat/plus_n_O.con"
-"cic:/matita/arithmetics/nat/plus_n_Sm.con"
-"cic:/matita/arithmetics/nat/plus_O_n.con"
-"cic:/matita/arithmetics/nat/commutative_plus.con"
-"cic:/matita/arithmetics/nat/associative_plus.con"
-"cic:/matita/arithmetics/nat/plus_plus_comm_23.con"
-"cic:/matita/arithmetics/nat/assoc_plus1.con"
-"cic:/matita/arithmetics/nat/injective_plus_l.con"
-"cic:/matita/arithmetics/nat/injective_plus_r.con"
-"cic:/matita/arithmetics/nat/nat_ind_plus.con"
-#
-"cic:/matita/arithmetics/nat/le.ind"
-"cic:/matita/arithmetics/nat/le_ind.con"
-"cic:/matita/arithmetics/nat/le_n_Sn.con"
-"cic:/matita/arithmetics/nat/le_O_n.con"
-"cic:/matita/arithmetics/nat/le_S_S.con"
-"cic:/matita/arithmetics/nat/le_or_ge.con"
-"cic:/matita/arithmetics/nat/le_S_S_to_le.con"
-"cic:/matita/arithmetics/nat/le_n_O_to_eq.con"
-"cic:/matita/arithmetics/nat/le_to_le_to_eq.con"
-"cic:/matita/arithmetics/nat/transitive_le.con"
-"cic:/matita/arithmetics/nat/decidable_le.con"
-%
-"cic:/matita/arithmetics/nat/monotonic_le_plus_l.con"
-"cic:/matita/arithmetics/nat/monotonic_le_plus_r.con"
-"cic:/matita/arithmetics/nat/le_plus_n_r.con"
-"cic:/matita/arithmetics/nat/le_plus_n.con"
-"cic:/matita/arithmetics/nat/le_plus_b.con"
-"cic:/matita/arithmetics/nat/le_plus_a.con"
-"cic:/matita/arithmetics/nat/le_plus.con"
-"cic:/matita/arithmetics/nat/plus_le_0.con"
-"cic:/matita/arithmetics/nat/le_plus_to_le_r.con"
-"cic:/matita/arithmetics/nat/le_plus_to_le.con"
-#
-"cic:/matita/arithmetics/nat/lt.con"
-"cic:/matita/arithmetics/nat/lt_O_S.con"
-"cic:/matita/arithmetics/nat/le_to_or_lt_eq.con"
-"cic:/matita/arithmetics/nat/eq_or_gt.con"
-"cic:/matita/arithmetics/nat/lt_or_ge.con"
-"cic:/matita/arithmetics/nat/not_le_to_lt.con"
-"cic:/matita/arithmetics/nat/lt_to_le_to_lt.con"
-"cic:/matita/arithmetics/nat/le_to_lt_to_lt.con"
-"cic:/matita/arithmetics/nat/lt_to_not_le.con"
-"cic:/matita/arithmetics/nat/lt_to_not_eq.con"
-"cic:/matita/arithmetics/nat/lt_to_le.con"
-"cic:/matita/arithmetics/nat/ltn_to_ltO.con"
-"cic:/matita/arithmetics/nat/transitive_lt.con"
-"cic:/matita/arithmetics/nat/nat_elim1.con"
-"cic:/matita/arithmetics/nat/f_ind.con"
-"cic:/matita/arithmetics/nat/f2_ind.con"
-"cic:/matita/arithmetics/nat/f3_ind.con"
-#
-"cic:/matita/arithmetics/nat/monotonic_lt_plus_l.con"
-"cic:/matita/arithmetics/nat/monotonic_lt_plus_r.con"
-"cic:/matita/arithmetics/nat/lt_plus_Sn_r.con"
-"cic:/matita/arithmetics/nat/lt_plus_to_lt_l.con"
-#
-"cic:/matita/arithmetics/nat/pred.con"
-#
-"cic:/matita/arithmetics/nat/pred_Sn.con"
-#
-"cic:/matita/arithmetics/nat/le_pred_n.con"
-"cic:/matita/arithmetics/nat/monotonic_pred.con"
-#
-"cic:/matita/arithmetics/nat/S_pred.con"
-#
-"cic:/matita/arithmetics/nat/minus.con"
-"cic:/matita/arithmetics/nat/minus_n_O.con"
-"cic:/matita/arithmetics/nat/eq_minus_S_pred.con"
-"cic:/matita/arithmetics/nat/minus_O_n.con"
-"cic:/matita/arithmetics/nat/minus_S_S.con"
-"cic:/matita/arithmetics/nat/minus_n_n.con"
-"cic:/matita/arithmetics/nat/minus_Sn_n.con"
-"cic:/matita/arithmetics/nat/minus_minus_comm.con"
-#
-"cic:/matita/arithmetics/nat/minus_plus_m_m.con"
-"cic:/matita/arithmetics/nat/minus_plus.con"
-"cic:/matita/arithmetics/nat/minus_plus_plus_l.con"
-"cic:/matita/arithmetics/nat/plus_minus_plus_plus_l.con"
-"cic:/matita/arithmetics/nat/plus_to_minus.con"
-"cic:/matita/arithmetics/nat/discr_plus_xy_minus_xz.con"
-"cic:/matita/arithmetics/nat/discr_minus_x_xy.con"
-#
-"cic:/matita/arithmetics/nat/minus_le.con"
-"cic:/matita/arithmetics/nat/inv_eq_minus_O.con"
-"cic:/matita/arithmetics/nat/monotonic_le_minus_l.con"
-"cic:/matita/arithmetics/nat/monotonic_le_minus_r.con"
-"cic:/matita/arithmetics/nat/eq_minus_O.con"
-"cic:/matita/arithmetics/nat/minus_Sn_m.con"
-"cic:/matita/arithmetics/nat/minus_minus_m_m.con"
-#
-"cic:/matita/arithmetics/nat/le_plus_minus_m_m.con"
-"cic:/matita/arithmetics/nat/le_plus_to_minus.con"
-"cic:/matita/arithmetics/nat/le_plus_to_minus_r.con"
-"cic:/matita/arithmetics/nat/le_inv_plus_l.con"
-"cic:/matita/arithmetics/nat/plus_minus_m_m.con"
-"cic:/matita/arithmetics/nat/le_minus_to_plus.con"
-"cic:/matita/arithmetics/nat/le_minus_to_plus_r.con"
-"cic:/matita/arithmetics/nat/plus_minus.con"
-"cic:/matita/arithmetics/nat/plus_minus_associative.con"
-"cic:/matita/arithmetics/nat/minus_minus_associative.con"
-"cic:/matita/arithmetics/nat/minus_minus.con"
-"cic:/matita/arithmetics/nat/minus_le_minus_minus_comm.con"
-#
-"cic:/matita/arithmetics/nat/minus_pred_pred.con"
-"cic:/matita/arithmetics/nat/monotonic_lt_minus_l.con"
-#
-"cic:/matita/arithmetics/nat/lt_plus_to_minus.con"
-"cic:/matita/arithmetics/nat/lt_plus_to_minus_r.con"
-"cic:/matita/arithmetics/nat/lt_inv_plus_l.con"
-"cic:/matita/arithmetics/nat/lt_minus_to_plus.con"
-"cic:/matita/arithmetics/nat/lt_minus_to_plus_r.con"
-#
-"cic:/matita/arithmetics/nat/max.con"
-"cic:/matita/arithmetics/nat/commutative_max.con"
-#
-"cic:/matita/arithmetics/nat/to_max.con"
-"cic:/matita/arithmetics/nat/le_maxl.con"
-"cic:/matita/arithmetics/nat/le_maxr.con"
-
-####################################
-
"cic:/matita/arithmetics/nat/nat_discr.con"
"cic:/matita/arithmetics/nat/times.con"
"cic:/matita/arithmetics/nat/times_n_1.con"
(*** le_or_ge *)
lemma nle_ge_dis (m) (n): ∨∨ m ≤ n | n ≤ m.
-#m #n @(nat_ind_succ_2 … m n) -m -n
+#m #n @(nat_ind_2_succ … m n) -m -n
[ /2 width=1 by or_introl/
| /2 width=1 by or_intror/
| #m #n * /3 width=2 by nle_succ_bi, or_introl, or_intror/
]
qed-.
-(* Basic inversions *********************************************************)
+(* Basic destructions *******************************************************)
-lemma nle_inv_succ_sn (m) (n): ↑m ≤ n → m ≤ n.
+lemma nle_des_succ_sn (m) (n): ↑m ≤ n → m ≤ n.
#m #n #H elim H -n /2 width=1 by nle_succ_dx/
qed-.
+(* Basic inversions *********************************************************)
+
(*** le_S_S_to_le *)
lemma nle_inv_succ_bi (m) (n): ↑m ≤ ↑n → m ≤ n.
#m #n @(insert_eq_0 … (↑n))
-#y * -y
-[ #H <(eq_inv_nsucc_bi … H) -m //
-| #y #Hy #H >(eq_inv_nsucc_bi … H) -n /2 width=1 by nle_inv_succ_sn/
+#x * -x
+[ #H >(eq_inv_nsucc_bi … H) -n //
+| #o #Ho #H >(eq_inv_nsucc_bi … H) -n
+ /2 width=1 by nle_des_succ_sn/
]
qed-.
#m @(insert_eq_0 … (𝟎))
#y * -y
[ #H destruct //
-| #y #_ #H elim (eq_inv_nzero_succ … H)
+| #y #_ #H elim (eq_inv_zero_nsucc … H)
]
qed-.
(* Advanced inversions ******************************************************)
+(*** le_plus_xSy_O_false *)
lemma nle_inv_succ_zero (m): ↑m ≤ 𝟎 → ⊥.
-/3 width=2 by nle_inv_zero_dx, eq_inv_nzero_succ/ qed-.
+/3 width=2 by nle_inv_zero_dx, eq_inv_zero_nsucc/ qed-.
lemma nle_inv_succ_sn_refl (m): ↑m ≤ m → ⊥.
#m @(nat_ind_succ … m) -m [| #m #IH ] #H
-[ /3 width=2 by nle_inv_zero_dx, eq_inv_nzero_succ/
+[ /3 width=2 by nle_inv_zero_dx, eq_inv_zero_nsucc/
| /3 width=1 by nle_inv_succ_bi/
]
qed-.
theorem nle_antisym (m) (n): m ≤ n → n ≤ m → m = n.
#m #n #H elim H -n //
#n #_ #IH #Hn
-lapply (nle_inv_succ_sn … Hn) #H
+lapply (nle_des_succ_sn … Hn) #H
lapply (IH H) -IH -H #H destruct
elim (nle_inv_succ_sn_refl … Hn)
qed-.
(∀n. Q (𝟎) (n)) →
(∀m,n. m ≤ n → Q m n → Q (↑m) (↑n)) →
∀m,n. m ≤ n → Q m n.
-#Q #IH1 #IH2 #m #n @(nat_ind_succ_2 … m n) -m -n //
+#Q #IH1 #IH2 #m #n @(nat_ind_2_succ … m n) -m -n //
[ #m #H elim (nle_inv_succ_zero … H)
| /4 width=1 by nle_inv_succ_bi/
]
(*** transitive_le *)
theorem nle_trans: Transitive … nle.
-#m #n #H elim H -n /3 width=1 by nle_inv_succ_sn/
+#m #n #H elim H -n /3 width=1 by nle_des_succ_sn/
qed-.
-(*** decidable_le *)
+(*** decidable_le le_dec *)
lemma nle_dec (m) (n): Decidable … (m ≤ n).
#m #n elim (nle_ge_dis m n) [ /2 width=1 by or_introl/ ]
#Hnm elim (eq_nat_dec m n) [ #H destruct /2 width=1 by nle_refl, or_introl/ ]
qed.
lemma nle_max_dx_refl_sn (m) (n): m ≤ (m ∨ n).
-#m #n @(nat_ind_succ_2 … m n) -m -n //
+#m #n @(nat_ind_2_succ … m n) -m -n //
#m #n #IH <nmax_succ_bi /2 width=1 by nle_succ_bi/
qed.
lemma nle_max_dx_refl_dx (m) (n): n ≤ (m ∨ n).
-#m #n @(nat_ind_succ_2 … n m) -m -n //
+#m #n @(nat_ind_2_succ … n m) -m -n //
#m #n #IH <nmax_succ_bi /2 width=1 by nle_succ_bi/
qed.
(*** inv_eq_minus_O *)
lemma nle_eq_minus_O (m) (n): 𝟎 = m - n → m ≤ n.
-#m #n @(nat_ind_succ_2 … m n) //
+#m #n @(nat_ind_2_succ … m n) //
/3 width=1 by nle_succ_bi/
qed.
#n #_ #IH /2 width=3 by nle_trans/
qed.
+(*** minus_le_trans_sn *)
+lemma nle_minus_sn (o) (m) (n): m ≤ n → m - o ≤ n.
+/2 width=3 by nle_trans/ qed.
+
(* Inversions with nminus ***************************************************)
(*** eq_minus_O *)
// qed.
(*** le_plus_to_minus *)
-lemma nle_minus_sn (o) (m) (n): m ≤ n + o → m - o ≤ n.
+lemma nle_minus_sn_sn (o) (m) (n): m ≤ n + o → m - o ≤ n.
+/2 width=1 by nle_minus_sn_bi/ qed.
+
+(*** le_plus_to_minus_comm *)
+lemma nle_minus_sn_dx (o) (m) (n): m ≤ o + n → m - o ≤ n.
/2 width=1 by nle_minus_sn_bi/ qed.
(*** le_plus_to_minus_r *)
-lemma nle_minus_dx (o) (m) (n): m + o ≤ n → m ≤ n - o.
+lemma nle_minus_dx_sn (o) (m) (n): m + o ≤ n → m ≤ n - o.
#o #m #n #H >(nminus_plus_sn_refl_sn m o)
/2 width=1 by nle_minus_sn_bi/
qed.
+(*** le_plus_to_minus_l *)
+lemma nle_minus_dx_dx (o) (m) (n): o + m ≤ n → m ≤ n - o.
+#o #m #n #H >(nminus_plus_sn_refl_dx m o)
+/2 width=1 by nle_minus_sn_bi/
+qed.
+
(*** le_inv_plus_l *)
lemma nle_minus_dx_full (o) (m) (n): m + o ≤ n → ∧∧ m ≤ n - o & o ≤ n.
-/3 width=3 by nle_minus_dx, nle_trans, conj/ qed-.
+/3 width=3 by nle_minus_dx_sn, nle_trans, conj/ qed-.
(* Inversions with nminus and nplus *****************************************)
<nplus_succ_sn //
qed-.
+(*** plus_minus_m_m_commutative *)
lemma nplus_minus_dx_refl_sn (m) (n): m ≤ n → n = m + (n - m).
#m #n <nplus_comm
/2 width=1 by nplus_minus_sn_refl_sn/
(* Destructions with nminus and nplus ***************************************)
-(*** plus_minus *)
+(*** plus_minus le_plus_minus_comm *)
lemma nminus_plus_comm_23 (o) (m) (n):
m ≤ n → n - m + o = n + o - m.
#o #m #n #H elim H -n //
/2 width=3 by nle_trans/
qed-.
-(*** plus_minus_associative *)
+(*** plus_minus_associative le_plus_minus *)
lemma nplus_minus_assoc (m1) (m2) (n):
n ≤ m2 → m1 + m2 - n = m1 + (m2 - n).
/2 width=1 by nminus_plus_comm_23/ qed-.
#m1 #m2 #m3 #Hm
>(nplus_minus_sn_refl_sn … Hm) in ⊢ (??%?); //
qed-.
+
+(*** plus_minus_minus_be *)
+lemma nplus_minus_be_be (m1) (m2) (m3):
+ m1 ≤ m2 → m2 ≤ m3 → (m3 - m2) + (m2 - m1) = m3 - m1.
+#m1 #m2 #m3 #Hm12 #Hm23
+>nminus_assoc // <nminus_minus_dx_refl_sn //
+qed-.
+
+(*** plus_to_minus_2 *)
+lemma nminus_plus_swap (m1) (m2) (n1) (n2):
+ m1 ≤ n1 → m2 ≤ n2 → m2+n1 = m1+n2 → n1-m1 = n2-m2.
+#m1 #m2 #n1 #n2 #H1 #H2 #H
+@nminus_plus_dx <nplus_minus_assoc //
+qed-.
(*** plus_le_0 *)
lemma nle_inv_plus_zero (m) (n): m + n ≤ 𝟎 → ∧∧ 𝟎 = m & 𝟎 = n.
-/3 width=1 by nle_inv_zero_dx, eq_inv_nzero_plus/ qed-.
+/3 width=1 by nle_inv_zero_dx, eq_inv_zero_nplus/ qed-.
(*** le_plus_to_le_r *)
lemma nle_inv_plus_bi_dx (o) (m) (n): n + o ≤ m + o → n ≤ m.
(*** le_plus_to_le *)
lemma nle_inv_plus_bi_sn (o) (m) (n): o + n ≤ o + m → n ≤ m.
/2 width=2 by nle_inv_plus_bi_dx/ qed-.
+
+(* Destructions with nplus **************************************************)
+
+(*** plus2_le_sn_sn *)
+lemma nplus_2_des_le_sn_sn (m1) (m2) (n1) (n2):
+ m1 + n1 = m2 + n2 → m1 ≤ m2 → n2 ≤ n1.
+#m1 #m2 #n1 #n2 #H #Hm
+lapply (nle_plus_bi_dx n1 … Hm) -Hm >H -H
+/2 width=2 by nle_inv_plus_bi_sn/
+qed-.
+
+(*** plus2_le_sn_dx *)
+lemma nplus_2_des_le_sn_dx (m1) (m2) (n1) (n2):
+ m1 + n1 = n2 + m2 → m1 ≤ m2 → n2 ≤ n1.
+#m1 #m2 #n1 #n2 <nplus_comm in ⊢ (???%→?);
+/2 width=4 by nplus_2_des_le_sn_sn/ qed-.
+
+(*** plus2_le_dx_sn *)
+lemma nplus_2_des_le_dx_sn (m1) (m2) (n1) (n2):
+ n1 + m1 = m2 + n2 → m1 ≤ m2 → n2 ≤ n1.
+#m1 #m2 #n1 #n2 <nplus_comm in ⊢ (??%?→?);
+/2 width=4 by nplus_2_des_le_sn_sn/ qed-.
+
+(*** plus2_le_dx_dx *)
+lemma nplus_2_des_le_dx_dx (m1) (m2) (n1) (n2):
+ n1 + m1 = n2 + m2 → m1 ≤ m2 → n2 ≤ n1.
+#m1 #m2 #n1 #n2 <nplus_comm in ⊢ (??%?→?);
+/2 width=4 by nplus_2_des_le_sn_dx/ qed-.
#m #n @(nat_ind_succ … m) -m //
/2 width=1 by nle_pred_bi/
qed-.
+
+(* Inversions with npred ****************************************************)
+
+(*** le_inv_S1 *)
+lemma nle_inv_succ_sn (m) (n):
+ ↑m ≤ n → ∧∧ m ≤ ↓n & n = ↑↓n.
+#m #n * -n
+[ /2 width=3 by nle_refl, conj/
+| #n #Hn /3 width=1 by nle_des_succ_sn, conj/
+]
+qed-.
(* *)
(**************************************************************************)
+include "ground/xoa/or_3.ma".
include "ground/arith/nat_le.ma".
(* STRICT ORDER FOR NON-NEGATIVE INTEGERS ***********************************)
lemma nlt_refl_succ (n): n < ↑n.
// qed.
+(*** lt_S *)
+lemma nlt_succ_dx (m) (n): m < n → m < ↑n.
+/2 width=1 by nle_succ_dx/ qed.
+
(*** lt_O_S *)
lemma nlt_zero_succ (m): 𝟎 < ↑m.
/2 width=1 by nle_succ_bi/ qed.
+(*** lt_S_S *)
lemma nlt_succ_bi (m) (n): m < n → ↑m < ↑n.
/2 width=1 by nle_succ_bi/ qed.
#H elim (nle_lt_eq_dis … H) -H /2 width=1 by nle_refl, or_introl, or_intror/
qed-.
+(*** lt_or_eq_or_gt *)
+lemma nlt_eq_gt_dis (m) (n): ∨∨ m < n | n = m | n < m.
+#m #n elim (nlt_ge_dis m n) /2 width=1 by or3_intro0/
+#H elim (nle_lt_eq_dis … H) -H /2 width=1 by or3_intro1, or3_intro2/
+qed-.
+
(*** not_le_to_lt *)
lemma le_false_nlt (m) (n): (n ≤ m → ⊥) → m < n.
#m #n elim (nlt_ge_dis m n) [ // ]
(* Basic inversions *********************************************************)
+(*** lt_S_S_to_lt *)
lemma nlt_inv_succ_bi (m) (n): ↑m < ↑n → m < n.
/2 width=1 by nle_inv_succ_bi/ qed-.
-(*** lt_to_not_le *)
+(*** lt_to_not_le lt_le_false *)
lemma nlt_ge_false (m) (n): m < n → n ≤ m → ⊥.
/3 width=4 by nle_inv_succ_sn_refl, nlt_le_trans/ qed-.
-(*** lt_to_not_eq *)
+(*** lt_to_not_eq lt_refl_false *)
lemma nlt_inv_refl (m): m < m → ⊥.
/2 width=4 by nlt_ge_false/ qed-.
+(*** lt_zero_false *)
lemma nlt_inv_zero_dx (m): m < 𝟎 → ⊥.
/2 width=4 by nlt_ge_false/ qed-.
(∀n. Q (𝟎) (↑n)) →
(∀m,n. m < n → Q m n → Q (↑m) (↑n)) →
∀m,n. m < n → Q m n.
-#Q #IH1 #IH2 #m #n @(nat_ind_succ_2 … n m) -m -n //
+#Q #IH1 #IH2 #m #n @(nat_ind_2_succ … n m) -m -n //
[ #m #H
elim (nlt_inv_zero_dx … H)
| /4 width=1 by nlt_inv_succ_bi/
<(nminus_succ_sn … Hom) //
qed.
+(*** monotonic_lt_minus_r *)
+lemma nlt_minus_dx_bi (o) (m) (n):
+ m < o -> m < n → o-n < o-m.
+#o #m #n #Ho #H
+lapply (nle_minus_dx_bi … o H) -H #H
+@(le_nlt_trans … H) -n
+@nlt_i >(nminus_succ_sn … Ho) //
+qed.
+
(* Destructions with nminus *************************************************)
(*** minus_pred_pred *)
lemma nminus_pred_bi (m) (n): 𝟎 < m → 𝟎 < n → n - m = ↓n - ↓m.
#m #n #Hm #Hn
->(nlt_inv_zero_sn … Hm) in ⊢ (??%?); -Hm
->(nlt_inv_zero_sn … Hn) in ⊢ (??%?); -Hn
+>(nlt_des_gen … Hm) in ⊢ (??%?); -Hm
+>(nlt_des_gen … Hn) in ⊢ (??%?); -Hn
//
qed-.
(*** lt_plus_to_minus *)
lemma nlt_minus_sn (o) (m) (n): m ≤ n → n < o + m → n - m < o.
#o #m #n #Hmn #Ho
-lapply (nle_minus_sn … Ho) -Ho
+lapply (nle_minus_sn_sn … Ho) -Ho
<nminus_succ_sn //
qed.
(*** lt_plus_to_minus_r *)
lemma nlt_minus_dx (o) (m) (n): m + o < n → m < n - o.
-/2 width=1 by nle_minus_dx/ qed.
+/2 width=1 by nle_minus_dx_sn/ qed.
(*** lt_inv_plus_l *)
lemma nlt_minus_dx_full (o) (m) (n): m + o < n → ∧∧ o < n & m < n - o.
@nlt_i >nplus_succ_dx /2 width=1 by nle_plus_bi_sn/
qed.
-(*** lt_plus_Sn_r *) (**)
-lemma lt_plus_Sn_r: ∀a,x,n. a < a + x + ↑n.
-/2 width=1/ qed-.
+lemma nlt_succ_plus_dx_refl_sn (m) (n): m < ↑(m + n).
+/2 width=1/ qed.
(* Inversions with nplus ****************************************************)
(* STRICT ORDER FOR NON-NEGATIVE INTEGERS ***********************************)
-(* Constructions with npred *************************************************)
+(* Destructions with npred **************************************************)
-lemma nlt_zero_sn (m): m = ↑↓m → 𝟎 < m.
-// qed.
+(*** S_pred lt_succ_pred lt_inv_O1 *)
+lemma nlt_des_gen (m) (n): m < n → n = ↑↓n.
+#m #n @(nat_ind_succ … n) -n //
+#H elim (nlt_inv_zero_dx … H)
+qed-.
(* Inversions with npred ****************************************************)
-(*** S_pred *)
-lemma nlt_inv_zero_sn (m): 𝟎 < m → m = ↑↓m.
-#m @(nat_ind_succ … m) -m //
-#H elim (nlt_inv_refl … H)
-qed-.
+(*** lt_inv_gen *)
+lemma nlt_inv_gen (m) (n): m < n → ∧∧ m ≤ ↓n & n = ↑↓n.
+/2 width=1 by nle_inv_succ_sn/ qed-.
+
+(*** lt_inv_S1 *)
+lemma nlt_inv_succ_sn (m) (n): ↑m < n → ∧∧ m < ↓n & n = ↑↓n.
+/2 width=1 by nle_inv_succ_sn/ qed-.
lemma nlt_inv_pred_dx (m) (n): m < ↓n → ↑m < n.
-#m #n #H >(nlt_inv_zero_sn n)
+#m #n #H >(nlt_des_gen (𝟎) n)
[ /2 width=1 by nlt_succ_bi/
| /3 width=3 by le_nlt_trans, nlt_le_trans/
]
qed-.
+
+(* Constructions with npred *************************************************)
+
+lemma nlt_zero_sn (n): n = ↑↓n → 𝟎 < n.
+// qed.
+
+(*** monotonic_lt_pred *)
+lemma nlt_pred_bi (m) (n): 𝟎 < m → m < n → ↓m < ↓n.
+#m #n #Hm #Hmn
+@nle_inv_succ_bi
+<(nlt_des_gen … Hm) -Hm
+<(nlt_des_gen … Hmn) //
+qed.
(* *)
(**************************************************************************)
-include "ground/notation/functions/zero_0.ma".
include "ground/arith/nat_succ_tri.ma".
(* MAXIMUM FOR NON-NEGATIVE INTEGERS ****************************************)
(*** max_SS *)
lemma nmax_succ_bi (n1) (n2): ↑(n1 ∨ n2) = (↑n1 ∨ ↑n2).
#n1 #n2
-@trans_eq [3: @ntri_succ_bi | skip ] (**) (* rewrite fails because δ-expansion gets in the way *)
+@trans_eq [3: @ntri_succ_bi | skip ] (**) (* rewrite fails because δ-expansion gets in the way *)
<ntri_f_tri //
qed.
(*** commutative_max *)
lemma nmax_comm: commutative … nmax.
-#m #n @(nat_ind_succ_2 … m n) -m -n //
+#m #n @(nat_ind_2_succ … m n) -m -n //
qed-.
(*** associative_max *)
lemma nmax_assoc: associative … nmax.
-#n1 #n2 @(nat_ind_succ_2 … n1 n2) -n1 -n2 //
+#n1 #n2 @(nat_ind_2_succ … n1 n2) -n1 -n2 //
#n1 #n2 #IH #n3 @(nat_ind_succ … n3) -n3 //
qed.
(*** max_inv_O3 *)
lemma nmax_inv_zero (n1) (n2): 𝟎 = (n1 ∨ n2) → ∧∧ 𝟎 = n1 & 𝟎 = n2.
-#n1 #n2 @(nat_ind_succ_2 … n1 n2) -n1 -n2 /2 width=1 by conj/
+#n1 #n2 @(nat_ind_2_succ … n1 n2) -n1 -n2 /2 width=1 by conj/
#n1 #n2 #_ <nmax_succ_bi #H
-elim (eq_inv_nzero_succ … H)
+elim (eq_inv_zero_nsucc … H)
qed-.
qed.
(*** minus_minus_comm *)
-lemma nminus_minus_comm (o) (m) (n): o - m - n = o - n - m.
+lemma nminus_comm (o) (m) (n): o - m - n = o - n - m.
#o #m #n @(nat_ind_succ … n) -n //
-qed-.
+qed.
+
+(*** minus_minus_comm3 *)
+lemma nminus_comm_231 (n) (m1) (m2) (m3):
+ n-m1-m2-m3 = n-m2-m3-m1.
+// qed.
#n #IH <nplus_succ_dx <nminus_succ_bi //
qed.
+(*** minus_plus_m_m *)
lemma nminus_plus_sn_refl_dx (m) (n): m = n + m - n.
#m #n <nplus_comm //
qed.
(*** minus_plus_plus_l *)
lemma nminus_plus_dx_bi (m) (n) (o): m - n = (m + o) - (n + o).
-#m #n #o <nminus_plus_assoc <nminus_minus_comm //
+#m #n #o <nminus_plus_assoc <nminus_comm //
qed.
-(*** plus_minus_plus_plus_l *) (**)
-lemma plus_minus_plus_plus_l: ∀z,x,y,h. z + (x + h) - (y + h) = z + x - y.
-// qed-.
-
(* Helper constructions with nplus ******************************************)
(*** plus_to_minus *)
m + o = m - n →
∨∨ ∧∧ 𝟎 = m & 𝟎 = o
| ∧∧ 𝟎 = n & 𝟎 = o.
-#m #n @(nat_ind_succ_2 … m n) -m -n
+#m #n @(nat_ind_2_succ … m n) -m -n
[ /3 width=1 by or_introl, conj/
| #m #o #Ho
lapply (eq_inv_nplus_bi_sn … (𝟎) Ho) -Ho
| #m #n #IH #o
<nminus_succ_bi >nplus_succ_shift #Ho
elim (IH … Ho) -IH -Ho * #_ #H
- elim (eq_inv_nzero_succ … H)
+ elim (eq_inv_zero_nsucc … H)
]
qed-.
(*** commutative_plus *)
lemma nplus_comm: commutative … nplus.
#m @(nat_ind_succ … m) -m //
-qed-.
+qed-. (**) (* gets in the way with auto *)
(*** associative_plus *)
lemma nplus_assoc: associative … nplus.
(* Basic inversions *********************************************************)
-lemma eq_inv_nzero_plus (m) (n): 𝟎 = m + n → ∧∧ 𝟎 = m & 𝟎 = n.
+(*** plus_inv_O3 zero_eq_plus *)
+lemma eq_inv_zero_nplus (m) (n): 𝟎 = m + n → ∧∧ 𝟎 = m & 𝟎 = n.
#m #n @(nat_ind_succ … n) -n
[ /2 width=1 by conj/
| #n #_ <nplus_succ_dx #H
- elim (eq_inv_nzero_succ … H)
+ elim (eq_inv_zero_nsucc … H)
]
qed-.
/2 width=2 by eq_inv_nplus_bi_dx/
qed-.
+(*** plus_xSy_x_false *)
+lemma succ_nplus_refl_sn (m) (n): m = ↑(m + n) → ⊥.
+#m @(nat_ind_succ … m) -m
+[ /2 width=2 by eq_inv_zero_nsucc/
+| #m #IH #n #H
+ @(IH n) /2 width=1 by eq_inv_nsucc_bi/
+]
+qed-.
+
+(*** discr_plus_xy_y *)
+lemma nplus_refl_dx (m) (n): n = m + n → 𝟎 = m.
+#m #n @(nat_ind_succ … n) -n //
+#n #IH /3 width=1 by eq_inv_nsucc_bi/
+qed-.
+
+(*** discr_plus_x_xy *)
+lemma nplus_refl_sn (m) (n): m = m + n → 𝟎 = n.
+#m #n <nplus_comm
+/2 width=2 by nplus_refl_dx/
+qed-.
+
(* Advanced eliminations ****************************************************)
(*** nat_ind_plus *)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||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 "ground/arith/nat_pred_succ.ma".
+include "ground/arith/nat_plus.ma".
+
+(* ADDITION FOR NON-NEGATIVE INTEGERS ***************************************)
+
+(* Inversions with npred ****************************************************)
+
+(*** plus_inv_S3_sn *)
+lemma eq_inv_succ_nplus_sn (o) (m) (n):
+ ↑o = m + n →
+ ∨∨ ∧∧ 𝟎 = m & n = ↑o
+ | ∧∧ m = ↑↓m & o = ↓m + n.
+#o #m @(nat_ind_succ … m) -m
+[ /3 width=1 by or_introl, conj/
+| #m #_ #n <nplus_succ_sn
+ /4 width=1 by eq_inv_nsucc_bi, or_intror, conj/
+]
+qed-.
+
+(*** plus_inv_S3_dx *)
+lemma eq_inv_succ_nplus_dx (o) (m) (n):
+ ↑o = m + n →
+ ∨∨ ∧∧ 𝟎 = n & m = ↑o
+ | ∧∧ n = ↑↓n & o = m + ↓n.
+#o #m #n @(nat_ind_succ … n) -n
+[ /3 width=1 by or_introl, conj/
+| #n #_ <nplus_succ_sn
+ /4 width=1 by eq_inv_nsucc_bi, or_intror, conj/
+]
+qed-.
(*** pred_O *)
lemma npred_zero: 𝟎 = ↓𝟎.
// qed.
+
+lemma npred_one: 𝟎 = ↓𝟏.
+// qed.
+
+lemma npred_psucc (p): ninj p = ↓↑p.
+// qed.
+
+(* Basic inversions *********************************************************)
+
+lemma npred_pnat_inv_refl (p): ninj p = ↓p → ⊥.
+*
+[ <npred_one #H destruct
+| #p /3 width=2 by psucc_inv_refl, eq_inv_ninj_bi/
+]
+qed-.
+
+(*** pred_inv_fix_sn *)
+lemma npred_inv_refl (n): n = ↓n → 𝟎 = n.
+* // #p #H elim (npred_pnat_inv_refl … H)
+qed-.
lemma npred_succ (n): n = ↓↑n.
* //
qed.
+
+(* Inversion with nsucc *****************************************************)
+
+(*** nat_split *)
+lemma nat_split (n): ∨∨ 𝟎 = n | n = ↑↓n.
+#n @(nat_ind_succ … n) -n
+/2 width=1 by or_introl, or_intror/
+qed-.
qed-.
(*** nat_elim2 *)
-lemma nat_ind_succ_2 (Q:relation2 …):
+lemma nat_ind_2_succ (Q:relation2 …):
(∀n. Q (𝟎) n) →
(∀m. Q (↑m) (𝟎)) →
(∀m,n. Q m n → Q (↑m) (↑n)) →
* [ <nsucc_zero | #p <nsucc_inj ] #H destruct
qed-.
-lemma eq_inv_nzero_succ (m): 𝟎 = ↑m → ⊥.
+lemma eq_inv_zero_nsucc (m): 𝟎 = ↑m → ⊥.
* [ <nsucc_zero | #p <nsucc_inj ] #H destruct
qed-.
+
+(*** succ_inv_refl_sn *)
+lemma nsucc_inv_refl (n): n = ↑n → ⊥.
+#n @(nat_ind_succ … n) -n
+[ /2 width=2 by eq_inv_zero_nsucc/
+| #n #IH #H /3 width=1 by eq_inv_nsucc_bi/
+]
+qed-.
#p #q #H destruct //
qed.
+lemma psucc_inv_refl (p): p = ↑p → ⊥.
+#p elim p -p
+[ #H destruct
+| #p #IH #H /3 width=1 by eq_inv_psucc_bi/
+]
+qed-.
+
(* Basic constructions ******************************************************)
lemma eq_pnat_dec (p1,p2:pnat): Decidable (p1 = p2).
lemma pplus_comm: commutative … pplus.
#p elim p -p //
-qed-.
+qed-. (**) (* gets in the way with auto *)
lemma pplus_assoc: associative … pplus.
#p #q #r elim r -r //
[ "nat_le ( ?≤? )" "nat_le_pred" "nat_le_plus" "nat_le_minus" "nat_le_minus_plus" "nat_le_max" * ]
[ "nat_max ( ?∨? )" * ]
[ "nat_minus ( ?-? )" "nat_minus_plus" * ]
- [ "nat_plus ( ?+? )" * ]
+ [ "nat_plus ( ?+? )" "nat_plus_pred" * ]
[ "nat_pred ( ↓? )" "nat_pred_succ" * ]
[ "nat_succ ( ↑? )" "nat_succ_iter" "nat_succ_tri" * ]
[ "nat ( 𝟎 )" "nat_iter ( ?^{?}? )" *"nat_tri" ]