renamed former Int.int_induct to Int.int_of_nat_induct, former Presburger.int_induct to Int.int_induct: is more conservative and more natural than the intermediate solution
authorhaftmann
Tue, 11 May 2010 08:36:02 +0200
changeset 36811 4ab4aa5bee1c
parent 36810 d9a51339746e
child 36813 75d837441aa6
renamed former Int.int_induct to Int.int_of_nat_induct, former Presburger.int_induct to Int.int_induct: is more conservative and more natural than the intermediate solution
NEWS
src/HOL/Int.thy
src/HOL/Library/Formal_Power_Series.thy
src/HOL/ex/Summation.thy
--- a/NEWS	Tue May 11 08:30:02 2010 +0200
+++ b/NEWS	Tue May 11 08:36:02 2010 +0200
@@ -140,8 +140,8 @@
 
 *** HOL ***
 
-* Theorem Presburger.int_induct has been renamed to Int.int_bidirectional_induct;
-theorem Int.int_induct is no longer shadowed.  INCOMPATIBILITY.
+* Theorem Int.int_induct renamed to Int.int_of_nat_induct and is
+no longer shadowed.  INCOMPATIBILITY.
 
 * Dropped theorem duplicate comp_arith; use semiring_norm instead.  INCOMPATIBILITY.
 
--- a/src/HOL/Int.thy	Tue May 11 08:30:02 2010 +0200
+++ b/src/HOL/Int.thy	Tue May 11 08:36:02 2010 +0200
@@ -559,7 +559,7 @@
 apply (blast dest: nat_0_le [THEN sym])
 done
 
-theorem int_induct [induct type: int, case_names nonneg neg]:
+theorem int_of_nat_induct [induct type: int, case_names nonneg neg]:
      "[|!! n. P (of_nat n \<Colon> int);  !!n. P (- (of_nat (Suc n))) |] ==> P z"
   by (cases z rule: int_cases) auto
 
@@ -1784,7 +1784,7 @@
 apply (rule step, simp+)
 done
 
-theorem int_bidirectional_induct [case_names base step1 step2]:
+theorem int_induct [case_names base step1 step2]:
   fixes k :: int
   assumes base: "P k"
     and step1: "\<And>i. k \<le> i \<Longrightarrow> P i \<Longrightarrow> P (i + 1)"
--- a/src/HOL/Library/Formal_Power_Series.thy	Tue May 11 08:30:02 2010 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy	Tue May 11 08:36:02 2010 +0200
@@ -402,7 +402,7 @@
 
 lemma number_of_fps_const: "(number_of k::('a::comm_ring_1) fps) = fps_const (of_int k)"
   
-proof(induct k rule: int_bidirectional_induct [where k=0])
+proof(induct k rule: int_induct [where k=0])
   case base thus ?case unfolding number_of_fps_def of_int_0 by simp
 next
   case (step1 i) thus ?case unfolding number_of_fps_def 
@@ -3214,7 +3214,7 @@
 
 lemma fps_number_of_fps_const: "number_of i = fps_const (number_of i :: 'a:: {comm_ring_1, number_ring})"
   apply (subst (2) number_of_eq)
-apply(rule int_bidirectional_induct[of _ 0])
+apply(rule int_induct [of _ 0])
 apply (simp_all add: number_of_fps_def)
 by (simp_all add: fps_const_add[symmetric] fps_const_minus[symmetric])
 
--- a/src/HOL/ex/Summation.thy	Tue May 11 08:30:02 2010 +0200
+++ b/src/HOL/ex/Summation.thy	Tue May 11 08:36:02 2010 +0200
@@ -42,7 +42,7 @@
   proof -
     fix k
     show "f k - g k = f 0 - g 0"
-      by (induct k rule: int_bidirectional_induct) (simp_all add: k_incr k_decr)
+      by (induct k rule: int_induct) (simp_all add: k_incr k_decr)
   qed
   then have "\<And>k. (plus (g 0 - f 0) \<circ> f) k = g k"
     by (simp add: algebra_simps)