# HG changeset patch # User wenzelm # Date 1622890672 -7200 # Node ID f143d0a4cb6af7781036a3220a863f6f3b300e07 # Parent 1c5dcba6925f129500a062c41acbcf523818c44d clarified examples; diff -r 1c5dcba6925f -r f143d0a4cb6a src/HOL/Examples/Sqrt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Examples/Sqrt.thy Sat Jun 05 12:57:52 2021 +0200 @@ -0,0 +1,104 @@ +(* Title: HOL/Examples/Sqrt.thy + Author: Makarius + Author: Tobias Nipkow, TU Muenchen +*) + +section \Square roots of primes are irrational\ + +theory Sqrt + imports Complex_Main "HOL-Computational_Algebra.Primes" +begin + +text \ + The square root of any prime number (including 2) is irrational. +\ + +theorem sqrt_prime_irrational: + fixes p :: nat + assumes "prime p" + shows "sqrt p \ \" +proof + 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) + 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 (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 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) + then obtain k where "m = p * k" .. + 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) + qed + then have "p dvd gcd m n" by simp + with \coprime m n\ have "p = 1" by simp + with p show False by simp +qed + +corollary sqrt_2_not_rat: "sqrt 2 \ \" + 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. +\ + +theorem + fixes p :: nat + assumes "prime p" + shows "sqrt p \ \" +proof + 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) + 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) + 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" by linarith + then have "p dvd m\<^sup>2" .. + 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 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) + 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\. +\ + +lemma "\a b::real. a \ \ \ b \ \ \ a powr b \ \" (is "\a b. ?P a b") +proof (cases "sqrt 2 powr sqrt 2 \ \") + case True + with sqrt_2_not_rat have "?P (sqrt 2) (sqrt 2)" by simp + then show ?thesis by blast +next + case False + with sqrt_2_not_rat powr_powr have "?P (sqrt 2 powr sqrt 2) (sqrt 2)" by simp + then show ?thesis by blast +qed + +end diff -r 1c5dcba6925f -r f143d0a4cb6a src/HOL/ROOT --- a/src/HOL/ROOT Sat Jun 05 12:45:00 2021 +0200 +++ b/src/HOL/ROOT Sat Jun 05 12:57:52 2021 +0200 @@ -20,7 +20,7 @@ Notable Examples in Isabelle/HOL. " sessions - "HOL-Library" + "HOL-Computational_Algebra" theories Adhoc_Overloading_Examples Ackermann @@ -36,6 +36,7 @@ Peirce Records Seq + Sqrt document_files "root.bib" "root.tex" @@ -706,7 +707,6 @@ Sketch_and_Explore Sorting_Algorithms_Examples Specifications_with_bundle_mixins - Sqrt Sqrt_Script Sudoku Sum_of_Powers diff -r 1c5dcba6925f -r f143d0a4cb6a src/HOL/ex/Sqrt.thy --- a/src/HOL/ex/Sqrt.thy Sat Jun 05 12:45:00 2021 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,104 +0,0 @@ -(* Title: HOL/ex/Sqrt.thy - Author: Makarius - Author: Tobias Nipkow, TU Muenchen -*) - -section \Square roots of primes are irrational\ - -theory Sqrt - imports Complex_Main "HOL-Computational_Algebra.Primes" -begin - -text \ - The square root of any prime number (including 2) is irrational. -\ - -theorem sqrt_prime_irrational: - fixes p :: nat - assumes "prime p" - shows "sqrt p \ \" -proof - 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) - 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 (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 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) - then obtain k where "m = p * k" .. - 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) - qed - then have "p dvd gcd m n" by simp - with \coprime m n\ have "p = 1" by simp - with p show False by simp -qed - -corollary sqrt_2_not_rat: "sqrt 2 \ \" - 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. -\ - -theorem - fixes p :: nat - assumes "prime p" - shows "sqrt p \ \" -proof - 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) - 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) - 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" by linarith - then have "p dvd m\<^sup>2" .. - 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 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) - 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\. -\ - -lemma "\a b::real. a \ \ \ b \ \ \ a powr b \ \" (is "\a b. ?P a b") -proof (cases "sqrt 2 powr sqrt 2 \ \") - case True - with sqrt_2_not_rat have "?P (sqrt 2) (sqrt 2)" by simp - then show ?thesis by blast -next - case False - with sqrt_2_not_rat powr_powr have "?P (sqrt 2 powr sqrt 2) (sqrt 2)" by simp - then show ?thesis by blast -qed - -end diff -r 1c5dcba6925f -r f143d0a4cb6a src/HOL/ex/Sqrt_Script.thy --- a/src/HOL/ex/Sqrt_Script.thy Sat Jun 05 12:45:00 2021 +0200 +++ b/src/HOL/ex/Sqrt_Script.thy Sat Jun 05 12:57:52 2021 +0200 @@ -5,14 +5,14 @@ section \Square roots of primes are irrational (script version)\ -theory Sqrt_Script -imports Complex_Main "HOL-Computational_Algebra.Primes" -begin +text \ + Contrast this linear Isabelle/Isar script with the more mathematical version + in \<^file>\~~/src/HOL/Examples/Sqrt.thy\ by Makarius Wenzel. +\ -text \ - \medskip Contrast this linear Isabelle/Isar script with Markus - Wenzel's more mathematical version. -\ +theory Sqrt_Script + imports Complex_Main "HOL-Computational_Algebra.Primes" +begin subsection \Preliminaries\