# HG changeset patch # User wenzelm # Date 1015433288 -3600 # Node ID 5765aa72afac56248b07014bb3a928562c710d75 # Parent 84e4ba7fb033a9a47b43bd1dc6e31295f06460d4 moved to Hyperreal-ex; diff -r 84e4ba7fb033 -r 5765aa72afac src/HOL/Real/ex/ROOT.ML --- a/src/HOL/Real/ex/ROOT.ML Wed Mar 06 17:47:51 2002 +0100 +++ b/src/HOL/Real/ex/ROOT.ML Wed Mar 06 17:48:08 2002 +0100 @@ -4,7 +4,4 @@ Miscellaneous HOL-Real Examples. *) -no_document use_thy "Primes"; -time_use_thy "Sqrt"; -time_use_thy "Sqrt_Script"; time_use_thy "BinEx"; diff -r 84e4ba7fb033 -r 5765aa72afac src/HOL/Real/ex/Sqrt.thy --- a/src/HOL/Real/ex/Sqrt.thy Wed Mar 06 17:47:51 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,152 +0,0 @@ -(* Title: HOL/Real/ex/Sqrt.thy - ID: $Id$ - Author: Markus Wenzel, TU Muenchen - License: GPL (GNU GENERAL PUBLIC LICENSE) -*) - -header {* Square roots of primes are irrational *} - -theory Sqrt = Primes + Real: - -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 add: power_two real_power_two 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_power_two) - then obtain k where "m = p * k" .. - with eq have "p * n\ = p\ * k\" by (auto simp add: power_two mult_ac) - with p have "n\ = p * k\" by (simp add: power_two) - hence "p dvd n\" .. - with p_prime show "p dvd n" by (rule prime_dvd_power_two) - 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::nat) ==> x \ \" -proof (rule sqrt_prime_irrational) - { - fix m :: nat assume dvd: "m dvd 2" - hence "m \ 2" by (simp add: dvd_imp_le) - moreover from dvd have "m \ 0" by (auto 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 add: power_two real_power_two 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_power_two) - then obtain k where "m = p * k" .. - with eq have "p * n\ = p\ * k\" by (auto simp add: power_two mult_ac) - with p have "n\ = p * k\" by (simp add: power_two) - hence "p dvd n\" .. - with p_prime have "p dvd n" by (rule prime_dvd_power_two) - 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 diff -r 84e4ba7fb033 -r 5765aa72afac src/HOL/Real/ex/Sqrt_Script.thy --- a/src/HOL/Real/ex/Sqrt_Script.thy Wed Mar 06 17:47:51 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,81 +0,0 @@ -(* Title: HOL/Real/ex/Sqrt_Script.thy - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 2001 University of Cambridge -*) - -header {* Square roots of primes are irrational *} - -theory Sqrt_Script = Primes + Real: - -text {* - \medskip Contrast this linear Isar script with Markus Wenzel's more - mathematical version. -*} - -subsection {* Preliminaries *} - -lemma prime_nonzero: "p \ prime \ p\0" -by (force simp add: prime_def) - -lemma prime_dvd_other_side: "\n*n = p*(k*k); p \ prime\ \ p dvd n" -apply (subgoal_tac "p dvd n*n", blast dest: prime_dvd_mult) -apply (rule_tac j="k*k" in dvd_mult_left, simp) -done - -lemma reduction: "\p \ prime; 0 < k; k*k = p*(j*j)\ \ k < p*j & 0 < j" -apply (rule ccontr) -apply (simp add: linorder_not_less) -apply (erule disjE) - apply (frule mult_le_mono, assumption) - apply auto -apply (force simp add: prime_def) -done - -lemma rearrange: "(j::nat) * (p*j) = k*k \ k*k = p*(j*j)" -by (simp add: mult_ac) - -lemma prime_not_square [rule_format]: - "p \ prime \ \k. 0 m*m \ p*(k*k)" -apply (induct_tac m rule: nat_less_induct) -apply clarify -apply (frule prime_dvd_other_side, assumption) -apply (erule dvdE) -apply (simp add: nat_mult_eq_cancel_disj prime_nonzero) -apply (blast dest: rearrange reduction) -done - - -subsection {* The set of rational numbers *} - -constdefs - rationals :: "real set" ("\") - "\ \ {x. \m n. n \ 0 \ \x\ = real (m::nat) / real (n::nat)}" - - -subsection {* Main theorem *} - -text {* - The square root of any prime number (including @{text 2}) is - irrational. -*} - -theorem prime_sqrt_irrational: "\p \ prime; x*x = real p; 0 \ x\ \ x \ \" -apply (simp add: rationals_def real_abs_def) -apply clarify -apply (erule_tac P="real m / real n * ?x = ?y" in rev_mp) -apply (simp del: real_of_nat_mult - add: real_divide_eq_eq prime_not_square - real_of_nat_mult [symmetric]) -done - -lemma two_is_prime: "2 \ prime" -apply (auto simp add: prime_def) -apply (case_tac m) -apply (auto dest!: dvd_imp_le) -apply arith -done - -lemmas two_sqrt_irrational = prime_sqrt_irrational [OF two_is_prime] - -end diff -r 84e4ba7fb033 -r 5765aa72afac src/HOL/Real/ex/document/root.tex --- a/src/HOL/Real/ex/document/root.tex Wed Mar 06 17:47:51 2002 +0100 +++ b/src/HOL/Real/ex/document/root.tex Wed Mar 06 17:48:08 2002 +0100 @@ -1,7 +1,8 @@ \documentclass[11pt,a4paper]{article} -\usepackage{isabelle,isabellesym,pdfsetup} \usepackage[latin1]{inputenc} +\usepackage{isabelle,isabellesym} +\usepackage{pdfsetup} \urlstyle{rm} \isabellestyle{it}