# HG changeset patch # User wenzelm # Date 1622889900 -7200 # Node ID 1c5dcba6925f129500a062c41acbcf523818c44d # Parent ce9529a616fd67077856db6456969b06e31933e7 tuned proofs; diff -r ce9529a616fd -r 1c5dcba6925f src/HOL/ex/Sqrt.thy --- a/src/HOL/ex/Sqrt.thy Sat Jun 05 12:29:57 2021 +0200 +++ b/src/HOL/ex/Sqrt.thy Sat Jun 05 12:45:00 2021 +0200 @@ -1,42 +1,46 @@ (* Title: HOL/ex/Sqrt.thy - Author: Markus Wenzel, Tobias Nipkow, TU Muenchen + Author: Makarius + Author: Tobias Nipkow, TU Muenchen *) section \Square roots of primes are irrational\ theory Sqrt -imports Complex_Main "HOL-Computational_Algebra.Primes" + imports Complex_Main "HOL-Computational_Algebra.Primes" begin -text \The square root of any prime number (including 2) is irrational.\ +text \ + The square root of any prime number (including 2) is irrational. +\ theorem sqrt_prime_irrational: - assumes "prime (p::nat)" + fixes p :: nat + assumes "prime p" shows "sqrt p \ \" proof - from \prime p\ have p: "1 < p" by (simp add: prime_nat_iff) + from \prime p\ have p: "p > 1" by (rule prime_gt_1_nat) assume "sqrt p \ \" - then obtain m n :: nat where - n: "n \ 0" and sqrt_rat: "\sqrt p\ = m / n" - and "coprime m n" by (rule Rats_abs_nat_div_natE) + then obtain m n :: nat + where n: "n \ 0" + and sqrt_rat: "\sqrt p\ = m / n" + and "coprime m n" by (rule Rats_abs_nat_div_natE) have eq: "m\<^sup>2 = p * n\<^sup>2" proof - from n and sqrt_rat have "m = \sqrt p\ * n" by simp - then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2" - by (auto simp add: power2_eq_square) + then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2" by (simp add: power_mult_distrib) also have "(sqrt p)\<^sup>2 = p" by simp also have "\ * n\<^sup>2 = p * n\<^sup>2" by simp - finally show ?thesis using of_nat_eq_iff by blast + finally show ?thesis by linarith qed have "p dvd m \ p dvd n" proof from eq have "p dvd m\<^sup>2" .. - with \prime p\ show "p dvd m" by (rule prime_dvd_power_nat) + with \prime p\ show "p dvd m" by (rule prime_dvd_power) 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 ac_simps) + with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by algebra 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) + with \prime p\ show "p dvd n" by (rule prime_dvd_power) qed then have "p dvd gcd m n" by simp with \coprime m n\ have "p = 1" by simp @@ -44,46 +48,47 @@ qed corollary sqrt_2_not_rat: "sqrt 2 \ \" - using sqrt_prime_irrational[of 2] by simp - - -subsection \Variations\ + using sqrt_prime_irrational [of 2] by simp text \ - Here is an alternative version of the main proof, using mostly - linear forward-reasoning. While this results in less top-down - structure, it is probably closer to proofs seen in mathematics. + Here is an alternative version of the main proof, using mostly linear + forward-reasoning. While this results in less top-down structure, it is + probably closer to proofs seen in mathematics. \ theorem - assumes "prime (p::nat)" + fixes p :: nat + assumes "prime p" shows "sqrt p \ \" proof - from \prime p\ have p: "1 < p" by (simp add: prime_nat_iff) + from \prime p\ have p: "p > 1" by (rule prime_gt_1_nat) assume "sqrt p \ \" - then obtain m n :: nat where - n: "n \ 0" and sqrt_rat: "\sqrt p\ = m / n" - and "coprime m n" by (rule Rats_abs_nat_div_natE) + then obtain m n :: nat + where n: "n \ 0" + and sqrt_rat: "\sqrt p\ = m / n" + and "coprime m n" by (rule Rats_abs_nat_div_natE) from n and sqrt_rat have "m = \sqrt p\ * n" by simp - then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2" - by (auto simp add: power2_eq_square) + then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2" by (auto simp add: power2_eq_square) also have "(sqrt p)\<^sup>2 = p" by simp also have "\ * n\<^sup>2 = p * n\<^sup>2" by simp - finally have eq: "m\<^sup>2 = p * n\<^sup>2" using of_nat_eq_iff by blast + finally have eq: "m\<^sup>2 = p * n\<^sup>2" by linarith then have "p dvd m\<^sup>2" .. - with \prime p\ have dvd_m: "p dvd m" by (rule prime_dvd_power_nat) + with \prime p\ have dvd_m: "p dvd m" by (rule prime_dvd_power) 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 ac_simps) + with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by algebra 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) + with \prime p\ have "p dvd n" by (rule prime_dvd_power) with dvd_m have "p dvd gcd m n" by (rule gcd_greatest) with \coprime m n\ have "p = 1" by simp with p show False by simp qed -text \Another old chestnut, which is a consequence of the irrationality of \<^term>\sqrt 2\.\ +text \ + Another old chestnut, which is a consequence of the irrationality of + \<^term>\sqrt 2\. +\ lemma "\a b::real. a \ \ \ b \ \ \ a powr b \ \" (is "\a b. ?P a b") proof (cases "sqrt 2 powr sqrt 2 \ \")