make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
authorhuffman
Mon, 23 Feb 2009 16:25:52 -0800
changeset 30079 293b896b9c25
parent 30078 beee83623cc9
child 30080 4cf42465b3da
make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
src/HOL/Arith_Tools.thy
src/HOL/Divides.thy
src/HOL/Groebner_Basis.thy
src/HOL/Int.thy
src/HOL/IntDiv.thy
src/HOL/List.thy
src/HOL/Nat.thy
src/HOL/NatBin.thy
src/HOL/Power.thy
src/HOL/Relation_Power.thy
src/HOL/SetInterval.thy
--- a/src/HOL/Arith_Tools.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/Arith_Tools.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -68,8 +68,9 @@
 apply (subst add_eq_if)
 apply (simp split add: nat.split
             del: nat_numeral_1_eq_1
-            add: numeral_1_eq_Suc_0 [symmetric] Let_def
-                 neg_imp_number_of_eq_0 neg_number_of_pred_iff_0)
+            add: nat_numeral_1_eq_1 [symmetric]
+                 numeral_1_eq_Suc_0 [symmetric]
+                 neg_number_of_pred_iff_0)
 done
 
 lemma nat_rec_number_of [simp]:
@@ -89,7 +90,8 @@
 apply (subst add_eq_if)
 apply (simp split add: nat.split
             del: nat_numeral_1_eq_1
-            add: numeral_1_eq_Suc_0 [symmetric] Let_def neg_imp_number_of_eq_0
+            add: nat_numeral_1_eq_1 [symmetric]
+                 numeral_1_eq_Suc_0 [symmetric]
                  neg_number_of_pred_iff_0)
 done
 
--- a/src/HOL/Divides.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/Divides.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -490,9 +490,9 @@
   from divmod_rel have divmod_m_n: "divmod_rel m n (m div n) (m mod n)" .
   with assms have m_div_n: "m div n \<ge> 1"
     by (cases "m div n") (auto simp add: divmod_rel_def)
-  from assms divmod_m_n have "divmod_rel (m - n) n (m div n - 1) (m mod n)"
+  from assms divmod_m_n have "divmod_rel (m - n) n (m div n - Suc 0) (m mod n)"
     by (cases "m div n") (auto simp add: divmod_rel_def)
-  with divmod_eq have "divmod (m - n) n = (m div n - 1, m mod n)" by simp
+  with divmod_eq have "divmod (m - n) n = (m div n - Suc 0, m mod n)" by simp
   moreover from divmod_div_mod have "divmod (m - n) n = ((m - n) div n, (m - n) mod n)" .
   ultimately have "m div n = Suc ((m - n) div n)"
     and "m mod n = (m - n) mod n" using m_div_n by simp_all
@@ -816,6 +816,9 @@
 lemma dvd_1_iff_1 [simp]: "(m dvd Suc 0) = (m = Suc 0)"
 by (simp add: dvd_def)
 
+lemma nat_dvd_1_iff_1 [simp]: "m dvd (1::nat) \<longleftrightarrow> m = 1"
+by (simp add: dvd_def)
+
 lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m = (n::nat)"
   unfolding dvd_def
   by (force dest: mult_eq_self_implies_10 simp add: mult_assoc mult_eq_1_iff)
--- a/src/HOL/Groebner_Basis.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/Groebner_Basis.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -147,7 +147,7 @@
 next show "pwr (mul x y) q = mul (pwr x q) (pwr y q)" by (rule pwr_mul)
 next show "pwr (pwr x p) q = pwr x (p * q)" by (rule pwr_pwr)
 next show "pwr x 0 = r1" using pwr_0 .
-next show "pwr x 1 = x" by (simp add: nat_number pwr_Suc pwr_0 mul_1 mul_c)
+next show "pwr x 1 = x" unfolding One_nat_def by (simp add: nat_number pwr_Suc pwr_0 mul_1 mul_c)
 next show "mul x (add y z) = add (mul x y) (mul x z)" using mul_d by simp
 next show "pwr x (Suc q) = mul x (pwr x q)" using pwr_Suc by simp
 next show "pwr x (2 * n) = mul (pwr x n) (pwr x n)" by (simp add: nat_number mul_pwr)
--- a/src/HOL/Int.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/Int.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -832,8 +832,8 @@
                              le_imp_0_less [THEN order_less_imp_le])  
 next
   case (neg n)
-  thus ?thesis by (simp del: of_nat_Suc of_nat_add
-    add: algebra_simps of_nat_1 [symmetric] of_nat_add [symmetric])
+  thus ?thesis by (simp del: of_nat_Suc of_nat_add of_nat_1
+    add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric])
 qed
 
 lemma bin_less_0_simps:
@@ -1165,8 +1165,8 @@
                              le_imp_0_less [THEN order_less_imp_le])  
 next
   case (neg n)
-  thus ?thesis by (simp del: of_nat_Suc of_nat_add
-    add: algebra_simps of_nat_1 [symmetric] of_nat_add [symmetric])
+  thus ?thesis by (simp del: of_nat_Suc of_nat_add of_nat_1
+    add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric])
 qed
 
 text {* Less-Than or Equals *}
@@ -1785,11 +1785,12 @@
 lemma int_val_lemma:
      "(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) -->  
       f 0 \<le> k --> k \<le> f n --> (\<exists>i \<le> n. f i = (k::int))"
+unfolding One_nat_def
 apply (induct n, simp)
 apply (intro strip)
 apply (erule impE, simp)
 apply (erule_tac x = n in allE, simp)
-apply (case_tac "k = f (n+1) ")
+apply (case_tac "k = f (Suc n)")
 apply force
 apply (erule impE)
  apply (simp add: abs_if split add: split_if_asm)
@@ -1803,6 +1804,7 @@
          f m \<le> k; k \<le> f n |] ==> ? i. m \<le> i & i \<le> n & f i = (k::int)"
 apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k 
        in int_val_lemma)
+unfolding One_nat_def
 apply simp
 apply (erule exE)
 apply (rule_tac x = "i+m" in exI, arith)
--- a/src/HOL/IntDiv.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/IntDiv.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -1228,7 +1228,7 @@
 text{*Suggested by Matthias Daum*}
 lemma int_power_div_base:
      "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)"
-apply (subgoal_tac "k ^ m = k ^ ((m - 1) + 1)")
+apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)")
  apply (erule ssubst)
  apply (simp only: power_add)
  apply simp_all
--- a/src/HOL/List.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/List.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -1438,10 +1438,10 @@
 apply (auto split:nat.split)
 done
 
-lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - 1)"
+lemma last_conv_nth: "xs\<noteq>[] \<Longrightarrow> last xs = xs!(length xs - Suc 0)"
 by(induct xs)(auto simp:neq_Nil_conv)
 
-lemma butlast_conv_take: "butlast xs = take (length xs - 1) xs"
+lemma butlast_conv_take: "butlast xs = take (length xs - Suc 0) xs"
 by (induct xs, simp, case_tac xs, simp_all)
 
 
@@ -1588,7 +1588,7 @@
 done
 
 lemma butlast_take:
-  "n <= length xs ==> butlast (take n xs) = take (n - 1) xs"
+  "n <= length xs ==> butlast (take n xs) = take (n - Suc 0) xs"
 by (simp add: butlast_conv_take min_max.inf_absorb1 min_max.inf_absorb2)
 
 lemma butlast_drop: "butlast (drop n xs) = drop n (butlast xs)"
@@ -1639,7 +1639,7 @@
 done
 
 lemma take_hd_drop:
-  "n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (n+1) xs"
+  "n < length xs \<Longrightarrow> take n xs @ [hd (drop n xs)] = take (Suc n) xs"
 apply(induct xs arbitrary: n)
 apply simp
 apply(simp add:drop_Cons split:nat.split)
@@ -2458,7 +2458,7 @@
 done
 
 lemma length_remove1:
-  "length(remove1 x xs) = (if x : set xs then length xs - 1 else length xs)"
+  "length(remove1 x xs) = (if x : set xs then length xs - Suc 0 else length xs)"
 apply (induct xs)
  apply (auto dest!:length_pos_if_in_set)
 done
--- a/src/HOL/Nat.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/Nat.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -196,8 +196,8 @@
 
 instance proof
   fix n m q :: nat
-  show "0 \<noteq> (1::nat)" by simp
-  show "1 * n = n" by simp
+  show "0 \<noteq> (1::nat)" unfolding One_nat_def by simp
+  show "1 * n = n" unfolding One_nat_def by simp
   show "n * m = m * n" by (induct n) simp_all
   show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib)
   show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib)
@@ -307,18 +307,24 @@
 lemmas nat_distrib =
   add_mult_distrib add_mult_distrib2 diff_mult_distrib diff_mult_distrib2
 
-lemma mult_eq_1_iff [simp]: "(m * n = Suc 0) = (m = 1 & n = 1)"
+lemma mult_eq_1_iff [simp]: "(m * n = Suc 0) = (m = Suc 0 & n = Suc 0)"
   apply (induct m)
    apply simp
   apply (induct n)
    apply auto
   done
 
-lemma one_eq_mult_iff [simp,noatp]: "(Suc 0 = m * n) = (m = 1 & n = 1)"
+lemma one_eq_mult_iff [simp,noatp]: "(Suc 0 = m * n) = (m = Suc 0 & n = Suc 0)"
   apply (rule trans)
   apply (rule_tac [2] mult_eq_1_iff, fastsimp)
   done
 
+lemma nat_mult_eq_1_iff [simp]: "m * n = (1::nat) \<longleftrightarrow> m = 1 \<and> n = 1"
+  unfolding One_nat_def by (rule mult_eq_1_iff)
+
+lemma nat_1_eq_mult_iff [simp]: "(1::nat) = m * n \<longleftrightarrow> m = 1 \<and> n = 1"
+  unfolding One_nat_def by (rule one_eq_mult_iff)
+
 lemma mult_cancel1 [simp]: "(k * m = k * n) = (m = n | (k = (0::nat)))"
 proof -
   have "k \<noteq> 0 \<Longrightarrow> k * m = k * n \<Longrightarrow> m = n"
@@ -465,11 +471,11 @@
 lemma less_Suc_eq: "(m < Suc n) = (m < n | m = n)"
   unfolding less_Suc_eq_le le_less ..
 
-lemma less_one [iff, noatp]: "(n < (1::nat)) = (n = 0)"
+lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)"
   by (simp add: less_Suc_eq)
 
-lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)"
-  by (simp add: less_Suc_eq)
+lemma less_one [iff, noatp]: "(n < (1::nat)) = (n = 0)"
+  unfolding One_nat_def by (rule less_Suc0)
 
 lemma Suc_mono: "m < n ==> Suc m < Suc n"
   by simp
@@ -800,6 +806,7 @@
   done
 
 lemma ex_least_nat_less: "\<not>P(0) \<Longrightarrow> P(n::nat) \<Longrightarrow> \<exists>k<n. (\<forall>i\<le>k. \<not>P i) & P(k+1)"
+  unfolding One_nat_def
   apply (cases n)
    apply blast
   apply (frule (1) ex_least_nat_le)
@@ -1089,7 +1096,7 @@
    apply simp_all
   done
 
-lemma one_le_mult_iff [simp]: "(Suc 0 \<le> m * n) = (1 \<le> m & 1 \<le> n)"
+lemma one_le_mult_iff [simp]: "(Suc 0 \<le> m * n) = (Suc 0 \<le> m & Suc 0 \<le> n)"
   apply (induct m)
    apply simp
   apply (case_tac n)
@@ -1125,7 +1132,7 @@
   by (cases m) (auto intro: le_add1)
 
 text {* Lemma for @{text gcd} *}
-lemma mult_eq_self_implies_10: "(m::nat) = m * n ==> n = 1 | m = 0"
+lemma mult_eq_self_implies_10: "(m::nat) = m * n ==> n = Suc 0 | m = 0"
   apply (drule sym)
   apply (rule disjCI)
   apply (rule nat_less_cases, erule_tac [2] _)
@@ -1164,7 +1171,7 @@
   | of_nat_Suc: "of_nat (Suc m) = 1 + of_nat m"
 
 lemma of_nat_1 [simp]: "of_nat 1 = 1"
-  by simp
+  unfolding One_nat_def by simp
 
 lemma of_nat_add [simp]: "of_nat (m + n) = of_nat m + of_nat n"
   by (induct m) (simp_all add: add_ac)
@@ -1276,7 +1283,7 @@
 end
 
 lemma of_nat_id [simp]: "of_nat n = n"
-  by (induct n) auto
+  by (induct n) (auto simp add: One_nat_def)
 
 lemma of_nat_eq_id [simp]: "of_nat = id"
   by (auto simp add: expand_fun_eq)
@@ -1381,7 +1388,7 @@
 apply(induct_tac k)
  apply simp
 apply(erule_tac x="m+n" in meta_allE)
-apply(erule_tac x="m+n+1" in meta_allE)
+apply(erule_tac x="Suc(m+n)" in meta_allE)
 apply simp
 done
 
--- a/src/HOL/NatBin.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/NatBin.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -442,19 +442,13 @@
 (* These two can be useful when m = number_of... *)
 
 lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))"
-apply (case_tac "m")
-apply (simp_all add: numerals)
-done
+  unfolding One_nat_def by (cases m) simp_all
 
 lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))"
-apply (case_tac "m")
-apply (simp_all add: numerals)
-done
+  unfolding One_nat_def by (cases m) simp_all
 
 lemma power_eq_if: "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))"
-apply (case_tac "m")
-apply (simp_all add: numerals)
-done
+  unfolding One_nat_def by (cases m) simp_all
 
 
 subsection{*Comparisons involving (0::nat) *}
--- a/src/HOL/Power.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/Power.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -31,7 +31,7 @@
   by (induct n) (simp_all add: power_Suc)
 
 lemma power_one_right [simp]: "(a::'a::recpower) ^ 1 = a"
-  by (simp add: power_Suc)
+  unfolding One_nat_def by (simp add: power_Suc)
 
 lemma power_commutes: "(a::'a::recpower) ^ n * a = a * a ^ n"
   by (induct n) (simp_all add: power_Suc mult_assoc)
@@ -366,8 +366,8 @@
   "of_nat (m ^ n) = (of_nat m::'a::{semiring_1,recpower}) ^ n"
 by (induct n, simp_all add: power_Suc of_nat_mult)
 
-lemma nat_one_le_power [simp]: "1 \<le> i ==> Suc 0 \<le> i^n"
-by (insert one_le_power [of i n], simp)
+lemma nat_one_le_power [simp]: "Suc 0 \<le> i ==> Suc 0 \<le> i^n"
+by (rule one_le_power [of i n, unfolded One_nat_def])
 
 lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)"
 by (induct "n", auto)
@@ -452,4 +452,3 @@
 *}
 
 end
-
--- a/src/HOL/Relation_Power.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/Relation_Power.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -61,16 +61,16 @@
 
 lemma funpow_swap1: "f((f^n) x) = (f^n)(f x)"
 proof -
-  have "f((f^n) x) = (f^(n+1)) x" by simp
+  have "f((f^n) x) = (f^(n+1)) x" unfolding One_nat_def by simp
   also have "\<dots>  = (f^n o f^1) x" by (simp only: funpow_add)
-  also have "\<dots> = (f^n)(f x)" by simp
+  also have "\<dots> = (f^n)(f x)" unfolding One_nat_def by simp
   finally show ?thesis .
 qed
 
 lemma rel_pow_1 [simp]:
   fixes R :: "('a*'a)set"
   shows "R^1 = R"
-  by simp
+  unfolding One_nat_def by simp
 
 lemma rel_pow_0_I: "(x,x) : R^0"
   by simp
--- a/src/HOL/SetInterval.thy	Mon Feb 23 13:55:36 2009 -0800
+++ b/src/HOL/SetInterval.thy	Mon Feb 23 16:25:52 2009 -0800
@@ -352,11 +352,11 @@
 
 corollary image_Suc_atLeastAtMost[simp]:
   "Suc ` {i..j} = {Suc i..Suc j}"
-using image_add_atLeastAtMost[where k=1] by simp
+using image_add_atLeastAtMost[where k="Suc 0"] by simp
 
 corollary image_Suc_atLeastLessThan[simp]:
   "Suc ` {i..<j} = {Suc i..<Suc j}"
-using image_add_atLeastLessThan[where k=1] by simp
+using image_add_atLeastLessThan[where k="Suc 0"] by simp
 
 lemma image_add_int_atLeastLessThan:
     "(%x. x + (l::int)) ` {0..<u-l} = {l..<u}"
@@ -556,7 +556,7 @@
 qed
 
 lemma card_less_Suc2: "0 \<notin> M \<Longrightarrow> card {k. Suc k \<in> M \<and> k < i} = card {k \<in> M. k < Suc i}"
-apply (rule card_bij_eq [of "Suc" _ _ "\<lambda>x. x - 1"])
+apply (rule card_bij_eq [of "Suc" _ _ "\<lambda>x. x - Suc 0"])
 apply simp
 apply fastsimp
 apply auto
@@ -803,7 +803,7 @@
 
 lemma setsum_head_upt_Suc:
   "m < n \<Longrightarrow> setsum f {m..<n} = f m + setsum f {Suc m..<n}"
-apply(insert setsum_head_Suc[of m "n - 1" f])
+apply(insert setsum_head_Suc[of m "n - Suc 0" f])
 apply (simp add: atLeastLessThanSuc_atLeastAtMost[symmetric] algebra_simps)
 done
 
@@ -835,11 +835,11 @@
 
 corollary setsum_shift_bounds_cl_Suc_ivl:
   "setsum f {Suc m..Suc n} = setsum (%i. f(Suc i)){m..n}"
-by (simp add:setsum_shift_bounds_cl_nat_ivl[where k=1,simplified])
+by (simp add:setsum_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified])
 
 corollary setsum_shift_bounds_Suc_ivl:
   "setsum f {Suc m..<Suc n} = setsum (%i. f(Suc i)){m..<n}"
-by (simp add:setsum_shift_bounds_nat_ivl[where k=1,simplified])
+by (simp add:setsum_shift_bounds_nat_ivl[where k="Suc 0", simplified])
 
 lemma setsum_shift_lb_Suc0_0:
   "f(0::nat) = (0::nat) \<Longrightarrow> setsum f {Suc 0..k} = setsum f {0..k}"
@@ -883,6 +883,7 @@
     by (rule setsum_addf)
   also from ngt1 have "\<dots> = ?n*a + (\<Sum>i\<in>{..<n}. ?I i*d)" by simp
   also from ngt1 have "\<dots> = (?n*a + d*(\<Sum>i\<in>{1..<n}. ?I i))"
+    unfolding One_nat_def
     by (simp add: setsum_right_distrib atLeast0LessThan[symmetric] setsum_shift_lb_Suc0_0_upt mult_ac)
   also have "(1+1)*\<dots> = (1+1)*?n*a + d*(1+1)*(\<Sum>i\<in>{1..<n}. ?I i)"
     by (simp add: left_distrib right_distrib)
@@ -890,7 +891,7 @@
     by (cases n) (auto simp: atLeastLessThanSuc_atLeastAtMost)
   also from ngt1
   have "(1+1)*?n*a + d*(1+1)*(\<Sum>i\<in>{1..n - 1}. ?I i) = ((1+1)*?n*a + d*?I (n - 1)*?I n)"
-    by (simp only: mult_ac gauss_sum [of "n - 1"])
+    by (simp only: mult_ac gauss_sum [of "n - 1"], unfold One_nat_def)
        (simp add:  mult_ac trans [OF add_commute of_nat_Suc [symmetric]])
   finally show ?thesis by (simp add: algebra_simps)
 next
@@ -906,7 +907,8 @@
     "((1::nat) + 1) * (\<Sum>i\<in>{..<n::nat}. a + of_nat(i)*d) =
     of_nat(n) * (a + (a + of_nat(n - 1)*d))"
     by (rule arith_series_general)
-  thus ?thesis by (auto simp add: of_nat_id)
+  thus ?thesis
+    unfolding One_nat_def by (auto simp add: of_nat_id)
 qed
 
 lemma arith_series_int: