ntheorem monotonic_pred: monotonic ? le pred.
#n; #m; #lenm; nelim lenm; /2/; nqed.
-ntheorem le_S_S_to_le : ∀n,m:nat. S n ≤ S m → n ≤ m.
+ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m.
/2/; nqed.
-(*
-nchange with (pred (S n) ≤ pred (S m));
-nelim leSS; apply le_n.apply (trans_le ? (pred n1)).assumption.
-apply le_pred_n.
-qed.
+ntheorem lt_S_S_to_lt: ∀n,m. S n < S m \to n < m.
+/2/; nqed.
-theorem lt_S_S_to_lt: \forall n,m.
- S n < S m \to n < m.
-intros. apply le_S_S_to_le. assumption.
-qed.
+ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
+/2/; nqed.
-theorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
-intros;
-unfold lt in H;
-apply (le_S_S ? ? H).
-qed.
+ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m.
+#n; #m; #Hlt; nelim Hlt;//; nqed.
-theorem leS_to_not_zero : \forall n,m:nat. S n \leq m \to not_zero m.
-intros.elim H.exact I.exact I.
-qed.
+(* lt vs. le *)
+ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O.
+#n; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed.
-(* not le *)
-theorem not_le_Sn_O: \forall n:nat. S n \nleq O.
-intros.unfold Not.simplify.intros.apply (leS_to_not_zero ? ? H).
-qed.
+ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m.
+/3/; nqed.
-theorem not_le_Sn_n: \forall n:nat. S n \nleq n.
-intros.elim n.apply not_le_Sn_O.unfold Not.simplify.intros.cut (S n1 \leq n1).
-apply H.assumption.
-apply le_S_S_to_le.assumption.
-qed.
+ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m.
+/3/; nqed.
+
+ntheorem decidable_le: ∀n,m. decidable (n≤m).
+napply nat_elim2; #n; /2/;
+#m; #dec; ncases dec;/3/; nqed.
+
+ntheorem decidable_lt: ∀n,m. decidable (n < m).
+#n; #m; napply decidable_le ; nqed.
+
+ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n.
+#n; nelim n; /2/; nqed.
+
+ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m.
+/2/; nqed.
+
+ntheorem not_le_to_lt: ∀n,m. n ≰ m → m < n.
+napply nat_elim2; #n;
+ ##[#abs; napply False_ind;/2/;
+ ##|/2/;
+ ##|#m;#Hind;#HnotleSS; napply lt_to_lt_S_S;/3/;
+ ##]
+nqed.
+
+ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n.
+#n; #m; #Hltnm; nelim Hltnm;/3/; nqed.
+
+ntheorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n.
+#n; #m; #Hnlt; napply lt_S_to_le;
+(* something strange here: /2/ fails *)
+napply not_le_to_lt; napply Hnlt; nqed.
+
+ntheorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n.
+/2/; nqed.
+
+(* lt and le trans *)
+
+ntheorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p.
+#n; #m; #p; #H; #H1; nelim H1; /2/; nqed.
+
+ntheorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p.
+#n; #m; #p; #H; nelim H; /3/; nqed.
+
+ntheorem lt_S_to_lt: ∀n,m. S n < m → n < m.
+/2/; nqed.
+
+ntheorem ltn_to_ltO: ∀n,m:nat. n < m → O < m.
+/2/; nqed.
+
+(*
+theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat.
+(S O) \lt n \to O \lt (pred n).
+intros.
+apply (ltn_to_ltO (pred (S O)) (pred n) ?).
+ apply (lt_pred (S O) n);
+ [ apply (lt_O_S O)
+ | assumption
+ ]
+qed. *)
+
+ntheorem lt_O_n_elim: ∀n:nat. O < n →
+ ∀P:nat → Prop.(∀m:nat.P (S m)) → P n.
+#n; nelim n; //; #abs; napply False_ind; /2/; nqed.
+(*
theorem lt_pred: \forall n,m.
O < n \to n < m \to pred n < pred m.
apply nat_elim2
].
qed.
-theorem le_to_le_pred:
- ∀n,m. n ≤ m → pred n ≤ pred m.
-intros 2;
-elim n;
-[ simplify;
- apply le_O_n
-| simplify;
- elim m in H1 ⊢ %;
- [ elim (not_le_Sn_O ? H1)
- | simplify;
- apply le_S_S_to_le;
- assumption
- ]
-].
-qed.
+*)
(* le to lt or eq *)
-theorem le_to_or_lt_eq : \forall n,m:nat.
-n \leq m \to n < m \lor n = m.
-intros.elim H.
-right.reflexivity.
-left.unfold lt.apply le_S_S.assumption.
-qed.
-
-theorem Not_lt_n_n: ∀n. n ≮ n.
-intro;
-unfold Not;
-intro;
-unfold lt in H;
-apply (not_le_Sn_n ? H).
-qed.
+ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m.
+#n; #m; #lenm; nelim lenm; /3/; nqed.
(* not eq *)
-theorem lt_to_not_eq : \forall n,m:nat. n<m \to n \neq m.
-unfold Not.intros.cut ((le (S n) m) \to False).
-apply Hcut.assumption.rewrite < H1.
-apply not_le_Sn_n.
-qed.
+ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m.
+/2/; nqed.
-(*not lt*)
-theorem eq_to_not_lt: \forall a,b:nat.
-a = b \to a \nlt b.
+(*not lt
+ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b.
intros.
unfold Not.
intros.
[ assumption
| reflexivity
]
-qed.
+qed.
theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n.
intros;
generalize in match (transitive_le ? ? ? H2 H1);
intro;
apply (not_le_Sn_n ? H3).
-qed.
-
-(* le vs. lt *)
-theorem lt_to_le : \forall n,m:nat. n<m \to n \leq m.
-simplify.intros.unfold lt in H.elim H.
-apply le_S. apply le_n.
-apply le_S. assumption.
-qed.
-
-theorem lt_S_to_le : \forall n,m:nat. n < S m \to n \leq m.
-simplify.intros.
-apply le_S_S_to_le.assumption.
-qed.
-
-theorem not_le_to_lt: \forall n,m:nat. n \nleq m \to m<n.
-intros 2.
-apply (nat_elim2 (\lambda n,m.n \nleq m \to m<n)).
-intros.apply (absurd (O \leq n1)).apply le_O_n.assumption.
-unfold Not.unfold lt.intros.apply le_S_S.apply le_O_n.
-unfold Not.unfold lt.intros.apply le_S_S.apply H.intros.apply H1.apply le_S_S.
-assumption.
-qed.
+qed. *)
-theorem lt_to_not_le: \forall n,m:nat. n<m \to m \nleq n.
-unfold Not.unfold lt.intros 3.elim H.
-apply (not_le_Sn_n n H1).
-apply H2.apply lt_to_le. apply H3.
-qed.
-
-theorem not_lt_to_le: \forall n,m:nat. Not (lt n m) \to le m n.
-simplify.intros.
-apply lt_S_to_le.
-apply not_le_to_lt.exact H.
-qed.
-
-theorem le_to_not_lt: \forall n,m:nat. le n m \to Not (lt m n).
-intros.unfold Not.unfold lt.
-apply lt_to_not_le.unfold lt.
-apply le_S_S.assumption.
-qed.
-
-theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
-intros;
-elim (le_to_or_lt_eq ? ? H1);
-[ assumption
-| elim (H H2)
-].
-qed.
+ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
+#n; #m; #Hneq; #Hle; ncases (le_to_or_lt_eq ?? Hle); //;
+#Heq; nelim (Hneq Heq); nqed.
(* le elimination *)
-theorem le_n_O_to_eq : \forall n:nat. n \leq O \to O=n.
-intro.elim n.reflexivity.
-apply False_ind.
-apply not_le_Sn_O;
-[2: apply H1 | skip].
-qed.
+ntheorem le_n_O_to_eq : ∀n:nat. n ≤ O → O=n.
+#n; ncases n; //; #a ; #abs; nelim (not_le_Sn_O ? abs); nqed.
-theorem le_n_O_elim: \forall n:nat.n \leq O \to \forall P: nat \to Prop.
-P O \to P n.
-intro.elim n.
-assumption.
-apply False_ind.
-apply (not_le_Sn_O ? H1).
-qed.
+ntheorem le_n_O_elim: ∀n:nat. n ≤ O → ∀P: nat →Prop. P O → P n.
+#n; ncases n; //; #a; #abs; nelim (not_le_Sn_O ? abs); nqed.
-theorem le_n_Sm_elim : \forall n,m:nat.n \leq S m \to
-\forall P:Prop. (S n \leq S m \to P) \to (n=S m \to P) \to P.
-intros 4.elim H.
-apply H2.reflexivity.
-apply H3. apply le_S_S. assumption.
-qed.
+ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m →
+∀P:Prop. (S n ≤ S m → P) → (n=S m → P) → P.
+#n; #m; #Hle; #P; nelim Hle; /3/; nqed.
(* le and eq *)
-lemma le_to_le_to_eq: \forall n,m. n \le m \to m \le n \to n = m.
-apply nat_elim2
- [intros.apply le_n_O_to_eq.assumption
- |intros.apply sym_eq.apply le_n_O_to_eq.assumption
- |intros.apply eq_f.apply H
- [apply le_S_S_to_le.assumption
- |apply le_S_S_to_le.assumption
- ]
- ]
-qed.
-(* lt and le trans *)
-theorem lt_O_S : \forall n:nat. O < S n.
-intro. unfold. apply le_S_S. apply le_O_n.
-qed.
-
-theorem lt_to_le_to_lt: \forall n,m,p:nat. lt n m \to le m p \to lt n p.
-intros.elim H1.
-assumption.unfold lt.apply le_S.assumption.
-qed.
-
-theorem le_to_lt_to_lt: \forall n,m,p:nat. le n m \to lt m p \to lt n p.
-intros 4.elim H.
-assumption.apply H2.unfold lt.
-apply lt_to_le.assumption.
-qed.
+ntheorem le_to_le_to_eq: ∀n,m. n ≤ m → m ≤ n → n = m.
+napply nat_elim2; /3/; nqed.
-theorem lt_S_to_lt: \forall n,m. S n < m \to n < m.
-intros.
-apply (trans_lt ? (S n))
- [apply le_n|assumption]
-qed.
-
-theorem ltn_to_ltO: \forall n,m:nat. lt n m \to lt O m.
-intros.apply (le_to_lt_to_lt O n).
-apply le_O_n.assumption.
-qed.
-
-theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat.
-(S O) \lt n \to O \lt (pred n).
-intros.
-apply (ltn_to_ltO (pred (S O)) (pred n) ?).
- apply (lt_pred (S O) n);
- [ apply (lt_O_S O)
- | assumption
- ]
-qed.
-
-theorem lt_O_n_elim: \forall n:nat. lt O n \to
-\forall P:nat\to Prop. (\forall m:nat.P (S m)) \to P n.
-intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
-apply H2.
-qed.
+ntheorem lt_O_S : \forall n:nat. O < S n.
+/2/; nqed.
+(*
(* other abstract properties *)
theorem antisymmetric_le : antisymmetric nat le.
unfold antisymmetric.intros 2.
apply antisym_le;
assumption.
qed.
-
-theorem decidable_le: \forall n,m:nat. decidable (n \leq m).
-intros.
-apply (nat_elim2 (\lambda n,m.decidable (n \leq m))).
-intros.unfold decidable.left.apply le_O_n.
-intros.unfold decidable.right.exact (not_le_Sn_O n1).
-intros 2.unfold decidable.intro.elim H.
-left.apply le_S_S.assumption.
-right.unfold Not.intro.apply H1.apply le_S_S_to_le.assumption.
-qed.
-
-theorem decidable_lt: \forall n,m:nat. decidable (n < m).
-intros.exact (decidable_le (S n) m).
-qed.
+*)
(* well founded induction principles *)
-theorem nat_elim1 : \forall n:nat.\forall P:nat \to Prop.
-(\forall m.(\forall p. (p \lt m) \to P p) \to P m) \to P n.
-intros.cut (\forall q:nat. q \le n \to P q).
-apply (Hcut n).apply le_n.
-elim n.apply (le_n_O_elim q H1).
-apply H.
-intros.apply False_ind.apply (not_le_Sn_O p H2).
-apply H.intros.apply H1.
-cut (p < S n1).
-apply lt_S_to_le.assumption.
-apply (lt_to_le_to_lt p q (S n1) H3 H2).
-qed.
+ntheorem nat_elim1 : ∀n:nat.∀P:nat → Prop.
+(∀m.(∀p. p < m → P p) → P m) → P n.
+#n; #P; #H;
+ncut (∀q:nat. q ≤ n → P q);/2/;
+nelim n;
+ ##[#q; #HleO; (* applica male *)
+ napply (le_n_O_elim ? HleO);
+ napply H; #p; #ltpO;
+ napply False_ind; /2/;
+ ##|#p; #Hind; #q; #HleS;
+ napply H; #a; #lta; napply Hind;
+ napply le_S_S_to_le;/2/;
+ ##]
+nqed.
(* some properties of functions *)
-
+(*
definition increasing \def \lambda f:nat \to nat.
\forall n:nat. f n < f (S n).
apply le_to_or_lt_eq.apply H6.
qed.
*)
+
+(******************* monotonicity ******************************)
+ntheorem monotonic_le_plus_r:
+∀n:nat.monotonic nat le (λm.n + m).
+#n; #a; #b; nelim n; nnormalize; //;
+#m; #H; #leab;napply le_S_S; /2/; nqed.
+
+ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m
+≝ monotonic_le_plus_r.
+
+ntheorem monotonic_le_plus_l:
+∀m:nat.monotonic nat le (λn.n + m).
+/2/; nqed.
+
+ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
+\def monotonic_le_plus_l.
+
+ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 \to m1 ≤ m2
+→ n1 + m1 ≤ n2 + m2.
+#n1; #n2; #m1; #m2; #len; #lem; napply transitive_le;
+/2/; nqed.
+
+ntheorem le_plus_n :∀n,m:nat. m ≤ n + m.
+/2/; nqed.
+
+ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n.
+/2/; nqed.
+
+ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n.
+//; nqed.
+
+ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
+#a; nelim a; /3/; nqed.
+
+(* times
+theorem monotonic_le_times_r:
+\forall n:nat.monotonic nat le (\lambda m. n * m).
+simplify.intros.elim n.
+simplify.apply le_O_n.
+simplify.apply le_plus.
+assumption.
+assumption.
+qed.
+
+theorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
+\def monotonic_le_times_r.
+
+theorem monotonic_le_times_l:
+\forall m:nat.monotonic nat le (\lambda n.n*m).
+simplify.intros.
+rewrite < sym_times.rewrite < (sym_times m).
+apply le_times_r.assumption.
+qed.
+
+theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
+\def monotonic_le_times_l.
+
+theorem le_times: \forall n1,n2,m1,m2:nat. n1 \le n2 \to m1 \le m2
+\to n1*m1 \le n2*m2.
+intros.
+apply (trans_le ? (n2*m1)).
+apply le_times_l.assumption.
+apply le_times_r.assumption.
+qed.
+
+theorem le_times_n: \forall n,m:nat.(S O) \le n \to m \le n*m.
+intros.elim H.simplify.
+elim (plus_n_O ?).apply le_n.
+simplify.rewrite < sym_plus.apply le_plus_n.
+qed.
+
+theorem le_times_to_le:
+\forall a,n,m. S O \le a \to a * n \le a * m \to n \le m.
+intro.
+apply nat_elim2;intros
+ [apply le_O_n
+ |apply False_ind.
+ rewrite < times_n_O in H1.
+ generalize in match H1.
+ apply (lt_O_n_elim ? H).
+ intros.
+ simplify in H2.
+ apply (le_to_not_lt ? ? H2).
+ apply lt_O_S
+ |apply le_S_S.
+ apply H
+ [assumption
+ |rewrite < times_n_Sm in H2.
+ rewrite < times_n_Sm in H2.
+ apply (le_plus_to_le a).
+ assumption
+ ]
+ ]
+qed.
+
+theorem le_S_times_SSO: \forall n,m.O < m \to
+n \le m \to S n \le (S(S O))*m.
+intros.
+simplify.
+rewrite > plus_n_O.
+simplify.rewrite > plus_n_Sm.
+apply le_plus
+ [assumption
+ |rewrite < plus_n_O.
+ assumption
+ ]
+qed.
+(*0 and times *)
+theorem O_lt_const_to_le_times_const: \forall a,c:nat.
+O \lt c \to a \le a*c.
+intros.
+rewrite > (times_n_SO a) in \vdash (? % ?).
+apply le_times
+[ apply le_n
+| assumption
+]
+qed. *)
\ No newline at end of file