src/HOL/ex/Sqrt.thy
changeset 31952 40501bb2d57c
parent 31712 6f8aa9aea693
child 32479 521cc9bf2958
--- a/src/HOL/ex/Sqrt.thy	Tue Jul 07 07:56:24 2009 +0200
+++ b/src/HOL/ex/Sqrt.thy	Tue Jul 07 17:39:51 2009 +0200
@@ -34,12 +34,12 @@
   have "p dvd m \<and> p dvd n"
   proof
     from eq have "p dvd m\<twosuperior>" ..
-    with `prime p` pos2 show "p dvd m" by (rule nat_prime_dvd_power)
+    with `prime p` pos2 show "p dvd m" by (rule prime_dvd_power_nat)
     then obtain k where "m = p * k" ..
     with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
     with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
     then have "p dvd n\<twosuperior>" ..
-    with `prime p` pos2 show "p dvd n" by (rule nat_prime_dvd_power)
+    with `prime p` pos2 show "p dvd n" by (rule prime_dvd_power_nat)
   qed
   then have "p dvd gcd m n" ..
   with gcd have "p dvd 1" by simp
@@ -48,7 +48,7 @@
 qed
 
 corollary "sqrt (real (2::nat)) \<notin> \<rat>"
-  by (rule sqrt_prime_irrational) (rule nat_two_is_prime)
+  by (rule sqrt_prime_irrational) (rule two_is_prime_nat)
 
 
 subsection {* Variations *}
@@ -75,13 +75,13 @@
   also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
   finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
   then have "p dvd m\<twosuperior>" ..
-  with `prime p` pos2 have dvd_m: "p dvd m" by (rule nat_prime_dvd_power)
+  with `prime p` pos2 have dvd_m: "p dvd m" by (rule prime_dvd_power_nat)
   then obtain k where "m = p * k" ..
   with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
   with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
   then have "p dvd n\<twosuperior>" ..
-  with `prime p` pos2 have "p dvd n" by (rule nat_prime_dvd_power)
-  with dvd_m have "p dvd gcd m n" by (rule nat_gcd_greatest)
+  with `prime p` pos2 have "p dvd n" by (rule prime_dvd_power_nat)
+  with dvd_m have "p dvd gcd m n" by (rule gcd_greatest_nat)
   with gcd have "p dvd 1" by simp
   then have "p \<le> 1" by (simp add: dvd_imp_le)
   with p show False by simp