src/HOL/ex/Sqrt.thy
changeset 57514 bdc2c6b40bf2
parent 53598 2bd8d459ebae
child 58889 5b7a9633cfa8
--- a/src/HOL/ex/Sqrt.thy	Sat Jul 05 10:09:01 2014 +0200
+++ b/src/HOL/ex/Sqrt.thy	Sat Jul 05 11:01:53 2014 +0200
@@ -33,7 +33,7 @@
     from eq have "p dvd m\<^sup>2" ..
     with `prime p` show "p dvd m" by (rule prime_dvd_power_nat)
     then obtain k where "m = p * k" ..
-    with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square mult_ac)
+    with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square ac_simps)
     with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
     then have "p dvd n\<^sup>2" ..
     with `prime p` show "p dvd n" by (rule prime_dvd_power_nat)
@@ -73,7 +73,7 @@
   then have "p dvd m\<^sup>2" ..
   with `prime p` have dvd_m: "p dvd m" by (rule prime_dvd_power_nat)
   then obtain k where "m = p * k" ..
-  with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square mult_ac)
+  with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square ac_simps)
   with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
   then have "p dvd n\<^sup>2" ..
   with `prime p` have "p dvd n" by (rule prime_dvd_power_nat)