]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user andrea
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Wed, 29 Feb 2012 11:31:22 +0000 (11:31 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Wed, 29 Feb 2012 11:31:22 +0000 (11:31 +0000)
weblib/arithmetics/nat.ma

index 24cfc5ef66d3e8d62c0eb5a6977134333b403cff..2260506329463e418b25706942c942bd3cf1c481 100644 (file)
@@ -552,7 +552,7 @@ monotonic_lt_plus_r. *)
 
 theorem monotonic_lt_plus_l: 
 ∀n:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6.\ 5a href="cic:/matita/basics/relations/monotonic.def(1)"\ 6monotonic\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/lt.def(1)"\ 6lt\ 5/a\ 6 (λm.m\ 5a title="natural plus" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6n).
-/\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/increasing_to_monotonic.def(4)"\ 6increasing_to_monotonic\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
+(* /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/increasing_to_monotonic.def(4)"\ 6increasing_to_monotonic\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ *) #n @\ 5a href="cic:/matita/arithmetics/nat/increasing_to_monotonic.def(4)"\ 6increasing_to_monotonic\ 5/a\ 6 // qed.
 
 (*
 variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def
@@ -945,7 +945,7 @@ qed.
 
 theorem monotonic_le_minus_r: 
 ∀p,q,n:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6. q \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 p → n\ 5a title="natural minus" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6\ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 n\ 5a title="natural minus" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6q.
-#p #q #n #lepq @\ 5a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(10)"\ 6le_plus_to_minus\ 5/a\ 6 \ 5span class="error" title="Parse error: illegal begin of statement"\ 6\ 5/span\ 6
+#p #q #n #lepq @\ 5a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(10)"\ 6le_plus_to_minus\ 5/a\ 6 
 @(\ 5a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"\ 6transitive_le\ 5/a\ 6 … (\ 5a href="cic:/matita/arithmetics/nat/le_plus_minus_m_m.def(9)"\ 6le_plus_minus_m_m\ 5/a\ 6 ? q)) /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/monotonic_le_plus_r.def(3)"\ 6monotonic_le_plus_r\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
 qed.
 
@@ -979,11 +979,11 @@ theorem plus_minus: ∀n,m,p. p ≤ m → (n+m)-p = n +(m-p).
 >associative_plus <plus_minus_m_m //
 qed.  *)
 
-theorem minus_minus: ∀n,m,p:nat. p ≤ m → m ≤ n →
-  p+(n-m) = n-(m-p).
+theorem minus_minus: ∀n,m,p:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6. p \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m → m \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 n →
+  p\ 5a title="natural plus" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6(n\ 5a title="natural minus" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6m) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 n\ 5a title="natural minus" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6(m\ 5a title="natural minus" href="cic:/fakeuri.def(1)"\ 6-\ 5/a\ 6p).
 #n #m #p #lepm #lemn
-@sym_eq @plus_to_minus <associative_plus <plus_minus_m_m //
-<commutative_plus <plus_minus_m_m //
+@\ 5a href="cic:/matita/basics/logic/sym_eq.def(2)"\ 6sym_eq\ 5/a\ 6 @\ 5a href="cic:/matita/arithmetics/nat/plus_to_minus.def(7)"\ 6plus_to_minus\ 5/a\ 6 <\ 5a href="cic:/matita/arithmetics/nat/associative_plus.def(4)"\ 6associative_plus\ 5/a\ 6 <\ 5a href="cic:/matita/arithmetics/nat/plus_minus_m_m.def(7)"\ 6plus_minus_m_m\ 5/a\ 6 //
+<\ 5a href="cic:/matita/arithmetics/nat/commutative_plus.def(5)"\ 6commutative_plus\ 5/a\ 6 <\ 5a href="cic:/matita/arithmetics/nat/plus_minus_m_m.def(7)"\ 6plus_minus_m_m\ 5/a\ 6 //
 qed.
 
 (*********************** boolean arithmetics ********************) 
@@ -991,20 +991,20 @@ include "basics/bool.ma".
 
 let rec eqb n m ≝ 
 match n with 
-  [ O ⇒ match m with [ O ⇒ true | S q ⇒ false
-  | S p ⇒ match m with [ O ⇒ false | S q ⇒ eqb p q]
+  [ O ⇒ match m with [ O ⇒ \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 | S q ⇒ \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
+  | S p ⇒ match m with [ O ⇒ \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 | S q ⇒ eqb p q]
   ].
 
-theorem eqb_elim : ∀ n,m:nat.∀ P:bool → Prop.
-(n=m → (P true)) → (n ≠ m → (P false)) → (P (eqb n m)). 
-@nat_elim2 
+theorem eqb_elim : ∀ n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6.∀ P:\ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 → Prop.
+(n\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6m → (P \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6)) → (n \ 5a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m → (P \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6)) → (P (\ 5a href="cic:/matita/arithmetics/nat/eqb.fix(0,0,1)"\ 6eqb\ 5/a\ 6 n m)). 
+@\ 5a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"\ 6nat_elim2\ 5/a\ 6 
   [#n (cases n) normalize /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5/span\ 6\ 5/span\ 6
   |normalize /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/sym_not_eq.def(4)"\ 6sym_not_eq\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
   |normalize /\ 5span class="autotactic"\ 64\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/not_eq_S.def(4)"\ 6not_eq_S\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6
   ] 
 qed.
 
-theorem eqb_n_n: ∀n. eqb n n = true.
+theorem eqb_n_n: ∀n. \ 5a href="cic:/matita/arithmetics/nat/eqb.fix(0,0,1)"\ 6eqb\ 5/a\ 6 n n \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6.
 #n (elim n) normalize // qed. 
 
 theorem eqb_true_to_eq: ∀n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6\ 5a href="cic:/matita/arithmetics/nat/eqb.fix(0,0,1)"\ 6eqb\ 5/a\ 6 n m \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 → n \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 m.
@@ -1013,7 +1013,7 @@ theorem eqb_true_to_eq: ∀n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,
 theorem eqb_false_to_not_eq: ∀n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6\ 5a href="cic:/matita/arithmetics/nat/eqb.fix(0,0,1)"\ 6eqb\ 5/a\ 6 n m \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 → n \ 5a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m.
 #n #m @(\ 5a href="cic:/matita/arithmetics/nat/eqb_elim.def(5)"\ 6eqb_elim\ 5/a\ 6 n m) /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/not_to_not.def(3)"\ 6not_to_not\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
 
-theorem eq_to_eqb_true: ∀n,m:nat.n = m → eqb n m = true.
+theorem eq_to_eqb_true: ∀n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6.n \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 m → \ 5a href="cic:/matita/arithmetics/nat/eqb.fix(0,0,1)"\ 6eqb\ 5/a\ 6 n m \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6.
 // qed.
 
 theorem not_eq_to_eqb_false: ∀n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6.
@@ -1022,15 +1022,15 @@ theorem not_eq_to_eqb_false: ∀n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind
 
 let rec leb n m ≝ 
 match n with 
-    [ O ⇒ true
+    [ O ⇒ \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6
     | (S p) ⇒
        match m with 
-        [ O ⇒ false
+        [ O ⇒ \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
              | (S q) ⇒ leb p q]].
 
-theorem leb_elim: ∀n,m:nat. ∀P:bool → Prop. 
-(n ≤ m → P true) → (n ≰ m → P false) → P (leb n m).
-@nat_elim2 normalize
+theorem leb_elim: ∀n,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6. ∀P:\ 5a href="cic:/matita/basics/bool/bool.ind(1,0,0)"\ 6bool\ 5/a\ 6 → Prop. 
+(n \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m → P \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6) → (n \ 5a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m → P \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6) → P (\ 5a href="cic:/matita/arithmetics/nat/leb.fix(0,0,1)"\ 6leb\ 5/a\ 6 n m).
+@\ 5a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"\ 6nat_elim2\ 5/a\ 6 normalize
   [/\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5/span\ 6\ 5/span\ 6/
   |/\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5/span\ 6\ 5/span\ 6/
   |#n #m #Hind #P #Pt #Pf @Hind
@@ -1051,7 +1051,7 @@ theorem le_to_leb_true: ∀n,m. n \ 5a title="natural 'less or equal to'" href="ci
 theorem not_le_to_leb_false: ∀n,m. n \ 5a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m → \ 5a href="cic:/matita/arithmetics/nat/leb.fix(0,0,1)"\ 6leb\ 5/a\ 6 n m \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
 #n #m @\ 5a href="cic:/matita/arithmetics/nat/leb_elim.def(6)"\ 6leb_elim\ 5/a\ 6 // #H #H1 @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/absurd.def(2)"\ 6absurd\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
 
-theorem lt_to_leb_false: ∀n,m. m < n → leb n m = false.
+theorem lt_to_leb_false: ∀n,m. m \ 5a title="natural 'less than'" href="cic:/fakeuri.def(1)"\ 6<\ 5/a\ 6 n → \ 5a href="cic:/matita/arithmetics/nat/leb.fix(0,0,1)"\ 6leb\ 5/a\ 6 n m \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
 /\ 5span class="autotactic"\ 63\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/lt_to_not_le.def(7)"\ 6lt_to_not_le\ 5/a\ 6\ 5a href="cic:/matita/arithmetics/nat/not_le_to_leb_false.def(7)"\ 6not_le_to_leb_false\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
 
 (* serve anche ltb? 
@@ -1095,11 +1095,11 @@ lemma commutative_min: \ 5a href="cic:/matita/basics/relations/commutative.def(1)"
 lemma le_minr: ∀i,n,m. i \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/min.def(2)"\ 6min\ 5/a\ 6 n m → i \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m.
 #i #n #m normalize @\ 5a href="cic:/matita/arithmetics/nat/leb_elim.def(6)"\ 6leb_elim\ 5/a\ 6 normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"\ 6transitive_le\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed. 
 
-lemma le_minl: ∀i,n,m. i ≤ min n m → i ≤ n.
+lemma le_minl: ∀i,n,m. i \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/min.def(2)"\ 6min\ 5/a\ 6 n m → i \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 n.
 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/le_minr.def(7)"\ 6le_minr\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
 
-lemma to_min: ∀i,n,m. i ≤ n → i ≤ m → i ≤ min n m.
-#i #n #m #lein #leim normalize (cases (leb n m)) 
+lemma to_min: ∀i,n,m. i \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 n → i \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 m → i \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/min.def(2)"\ 6min\ 5/a\ 6 n m.
+#i #n #m #lein #leim normalize (cases (\ 5a href="cic:/matita/arithmetics/nat/leb.fix(0,0,1)"\ 6leb\ 5/a\ 6 n m)) 
 normalize // qed.
 
 lemma commutative_max: \ 5a href="cic:/matita/basics/relations/commutative.def(1)"\ 6commutative\ 5/a\ 6 ? \ 5a href="cic:/matita/arithmetics/nat/max.def(2)"\ 6max\ 5/a\ 6.
@@ -1111,7 +1111,7 @@ lemma commutative_max: \ 5a href="cic:/matita/basics/relations/commutative.def(1)"
 lemma le_maxl: ∀i,n,m. \ 5a href="cic:/matita/arithmetics/nat/max.def(2)"\ 6max\ 5/a\ 6 n m \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i → n \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i.
 #i #n #m normalize @\ 5a href="cic:/matita/arithmetics/nat/leb_elim.def(6)"\ 6leb_elim\ 5/a\ 6 normalize /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"\ 6transitive_le\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed. 
 
-lemma le_maxr: ∀i,n,m. max n m ≤ i → m ≤ i.
+lemma le_maxr: ∀i,n,m. \ 5a href="cic:/matita/arithmetics/nat/max.def(2)"\ 6max\ 5/a\ 6 n m \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i → m \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i.
 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/arithmetics/nat/le_maxl.def(7)"\ 6le_maxl\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/ qed.
 
 lemma to_max: ∀i,n,m. n \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i → m \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i → \ 5a href="cic:/matita/arithmetics/nat/max.def(2)"\ 6max\ 5/a\ 6 n m \ 5a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i.