# HG changeset patch # User wenzelm # Date 1001609020 -7200 # Node ID 3ccea743e5e71053e9a39c3c6e0dc120852673b9 # Parent ab08d61966b10d52b760a17d4fde9367d9c29671 Square roots of primes are irrational; diff -r ab08d61966b1 -r 3ccea743e5e7 src/HOL/Real/ex/Sqrt_Irrational.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Real/ex/Sqrt_Irrational.thy Thu Sep 27 18:43:40 2001 +0200 @@ -0,0 +1,159 @@ +(* Title: HOL/Real/ex/Sqrt_Irrational.thy + ID: $Id$ + Author: Markus Wenzel, TU Muenchen + License: GPL (GNU GENERAL PUBLIC LICENSE) +*) + +header {* Square roots of primes are irrational *} + +theory Sqrt_Irrational = Real + Primes: + +syntax (xsymbols) (* FIXME move to main HOL (!?) *) + "_square" :: "'a \ 'a" ("(_\)" [1000] 999) +syntax (HTML output) + "_square" :: "'a \ 'a" ("(_\)" [1000] 999) +syntax (output) + "_square" :: "'a \ 'a" ("(_^2)" [1000] 999) +translations + "x\" \ "x^2" + + +subsection {* The set of rational numbers *} + +constdefs + rationals :: "real set" ("\") + "\ \ {x. \m n. n \ 0 \ \x\ = real (m::nat) / real (n::nat)}" + +theorem rationals_rep: "x \ \ \ + \m n. n \ 0 \ \x\ = real m / real n \ gcd (m, n) = 1" +proof - + assume "x \ \" + then obtain m n :: nat where n: "n \ 0" and x_rat: "\x\ = real m / real n" + by (unfold rationals_def) blast + let ?gcd = "gcd (m, n)" + from n have gcd: "?gcd \ 0" by (simp add: gcd_zero) + let ?k = "m div ?gcd" + let ?l = "n div ?gcd" + let ?gcd' = "gcd (?k, ?l)" + have "?gcd dvd m" .. hence gcd_k: "?gcd * ?k = m" + by (rule dvd_mult_div_cancel) + have "?gcd dvd n" .. hence gcd_l: "?gcd * ?l = n" + by (rule dvd_mult_div_cancel) + + from n gcd_l have "?l \ 0" + by (auto iff del: neq0_conv) + moreover + have "\x\ = real ?k / real ?l" + proof - + from gcd have "real ?k / real ?l = real (?gcd * ?k) / real (?gcd * ?l)" + by (simp add: real_mult_div_cancel1) + also from gcd_k gcd_l have "\ = real m / real n" by simp + also from x_rat have "\ = \x\" .. + finally show ?thesis .. + qed + moreover + have "?gcd' = 1" + proof - + have "?gcd * ?gcd' = gcd (?gcd * ?k, ?gcd * ?l)" + by (rule gcd_mult_distrib2) + with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp + with gcd show ?thesis by simp + qed + ultimately show ?thesis by blast +qed + +lemma [elim?]: "r \ \ \ + (\m n. n \ 0 \ \r\ = real m / real n \ gcd (m, n) = 1 \ C) \ C" + by (insert rationals_rep) blast + + +subsection {* Main theorem *} + +text {* + The square root of any prime number (including @{text 2}) is + irrational. +*} + +theorem sqrt_prime_irrational: "x\ = real p \ p \ prime \ x \ \" +proof + assume x_sqrt: "x\ = real p" + assume p_prime: "p \ prime" + hence p: "1 < p" by (simp add: prime_def) + assume "x \ \" + then obtain m n where + n: "n \ 0" and x_rat: "\x\ = real m / real n" and gcd: "gcd (m, n) = 1" .. + have eq: "m\ = p * n\" + proof - + from n x_rat have "real m = \x\ * real n" by simp + hence "real (m\) = x\ * real (n\)" by (simp split: abs_split) + also from x_sqrt have "\ = real (p * n\)" by simp + finally show ?thesis .. + qed + have "p dvd m \ p dvd n" + proof + from eq have "p dvd m\" .. + with p_prime show "p dvd m" by (rule prime_dvd_square) + then obtain k where "m = p * k" .. + with eq have "p * n\ = p\ * k\" by (auto simp add: mult_ac) + with p have "n\ = p * k\" by simp + hence "p dvd n\" .. + with p_prime show "p dvd n" by (rule prime_dvd_square) + qed + hence "p dvd gcd (m, n)" .. + with gcd have "p dvd 1" by simp + hence "p \ 1" by (simp add: dvd_imp_le) + with p show False by simp +qed + + +subsection {* Variations *} + +text {* + Just for the record: we instantiate the main theorem for the + specific prime number @{text 2} (real mathematicians would never do + this formally :-). +*} + +theorem "x\ = real 2 \ x \ \" +proof (rule sqrt_prime_irrational) + { + fix m assume dvd: "m dvd 2" + hence "m \ 2" by (simp add: dvd_imp_le) + moreover from dvd have "m \ 0" by (auto dest: dvd_0_left iff del: neq0_conv) + ultimately have "m = 1 \ m = 2" by arith + } + thus "2 \ prime" by (simp add: prime_def) +qed + +text {* + \medskip 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 "x\ = real p \ p \ prime \ x \ \" +proof + assume x_sqrt: "x\ = real p" + assume p_prime: "p \ prime" + hence p: "1 < p" by (simp add: prime_def) + assume "x \ \" + then obtain m n where + n: "n \ 0" and x_rat: "\x\ = real m / real n" and gcd: "gcd (m, n) = 1" .. + from n x_rat have "real m = \x\ * real n" by simp + hence "real (m\) = x\ * real (n\)" by (simp split: abs_split) + also from x_sqrt have "\ = real (p * n\)" by simp + finally have eq: "m\ = p * n\" .. + hence "p dvd m\" .. + with p_prime have dvd_m: "p dvd m" by (rule prime_dvd_square) + then obtain k where "m = p * k" .. + with eq have "p * n\ = p\ * k\" by (auto simp add: mult_ac) + with p have "n\ = p * k\" by simp + hence "p dvd n\" .. + with p_prime have "p dvd n" by (rule prime_dvd_square) + with dvd_m have "p dvd gcd (m, n)" by (rule gcd_greatest) + with gcd have "p dvd 1" by simp + hence "p \ 1" by (simp add: dvd_imp_le) + with p show False by simp +qed + +end