aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar David Aspinall <da@inf.ed.ac.uk>2004-04-16 14:30:13 +0000
committerGravatar David Aspinall <da@inf.ed.ac.uk>2004-04-16 14:30:13 +0000
commit2b17630a04cc510ef6c504614ab59e99071c5cce (patch)
tree35646939b64142d48e323dd1686fb74fbd67f3dd
parent98b419b104a9c7b0a12ee2f00248ef983117a725 (diff)
New files.
-rw-r--r--isar/Root2_Isar.thy151
-rw-r--r--isar/Root2_Tactic.thy81
2 files changed, 232 insertions, 0 deletions
diff --git a/isar/Root2_Isar.thy b/isar/Root2_Isar.thy
new file mode 100644
index 00000000..d09cc0d4
--- /dev/null
+++ b/isar/Root2_Isar.thy
@@ -0,0 +1,151 @@
+(* Example proof by Markus Wenzel; see http://www.cs.kun.nl/~freek/comparison/
+ Taken from Isabelle2004 distribution. *)
+
+
+(* Title: HOL/Hyperreal/ex/Sqrt.thy
+ ID: Id: Sqrt.thy,v 1.4 2004/01/12 15:51:47 paulson Exp
+ Author: Markus Wenzel, TU Muenchen
+ License: GPL (GNU GENERAL PUBLIC LICENSE)
+*)
+
+header {* Square roots of primes are irrational *}
+
+theory Root2_Isar = Primes + Complex_Main:
+
+subsection {* Preliminaries *}
+
+text {*
+ The set of rational numbers, including the key representation
+ theorem.
+*}
+
+constdefs
+ rationals :: "real set" ("\<rat>")
+ "\<rat> \<equiv> {x. \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real (m::nat) / real (n::nat)}"
+
+theorem rationals_rep: "x \<in> \<rat> \<Longrightarrow>
+ \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real m / real n \<and> gcd (m, n) = 1"
+proof -
+ assume "x \<in> \<rat>"
+ then obtain m n :: nat where
+ n: "n \<noteq> 0" and x_rat: "\<bar>x\<bar> = real m / real n"
+ by (unfold rationals_def) blast
+ let ?gcd = "gcd (m, n)"
+ from n have gcd: "?gcd \<noteq> 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" .. then have gcd_k: "?gcd * ?k = m"
+ by (rule dvd_mult_div_cancel)
+ have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
+ by (rule dvd_mult_div_cancel)
+
+ from n and gcd_l have "?l \<noteq> 0"
+ by (auto iff del: neq0_conv)
+ moreover
+ have "\<bar>x\<bar> = real ?k / real ?l"
+ proof -
+ from gcd have "real ?k / real ?l =
+ real (?gcd * ?k) / real (?gcd * ?l)"
+ by (simp add: mult_divide_cancel_left)
+ also from gcd_k and gcd_l have "\<dots> = real m / real n" by simp
+ also from x_rat have "\<dots> = \<bar>x\<bar>" ..
+ 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 \<in> \<rat> \<Longrightarrow>
+ (\<And>m n. n \<noteq> 0 \<Longrightarrow> \<bar>r\<bar> = real m / real n \<Longrightarrow> gcd (m, n) = 1 \<Longrightarrow> C)
+ \<Longrightarrow> C"
+ using rationals_rep by blast
+
+
+subsection {* Main theorem *}
+
+text {*
+ The square root of any prime number (including @{text 2}) is
+ irrational.
+*}
+
+theorem sqrt_prime_irrational: "p \<in> prime \<Longrightarrow> sqrt (real p) \<notin> \<rat>"
+proof
+ assume p_prime: "p \<in> prime"
+ then have p: "1 < p" by (simp add: prime_def)
+ assume "sqrt (real p) \<in> \<rat>"
+ then obtain m n where
+ n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
+ and gcd: "gcd (m, n) = 1" ..
+ have eq: "m\<twosuperior> = p * n\<twosuperior>"
+ proof -
+ from n and sqrt_rat have "real m = \<bar>sqrt (real p)\<bar> * real n" by simp
+ then have "real (m\<twosuperior>) = (sqrt (real p))\<twosuperior> * real (n\<twosuperior>)"
+ by (auto simp add: power2_eq_square)
+ also have "(sqrt (real p))\<twosuperior> = real p" by simp
+ also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
+ finally show ?thesis ..
+ qed
+ have "p dvd m \<and> p dvd n"
+ proof
+ from eq have "p dvd m\<twosuperior>" ..
+ 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\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
+ with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
+ then have "p dvd n\<twosuperior>" ..
+ with p_prime show "p dvd n" by (rule prime_dvd_power_two)
+ qed
+ then have "p dvd gcd (m, n)" ..
+ with gcd have "p dvd 1" by simp
+ then have "p \<le> 1" by (simp add: dvd_imp_le)
+ with p show False by simp
+qed
+
+corollary "sqrt (real (2::nat)) \<notin> \<rat>"
+ by (rule sqrt_prime_irrational) (rule two_is_prime)
+
+
+subsection {* Variations *}
+
+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 "p \<in> prime \<Longrightarrow> sqrt (real p) \<notin> \<rat>"
+proof
+ assume p_prime: "p \<in> prime"
+ then have p: "1 < p" by (simp add: prime_def)
+ assume "sqrt (real p) \<in> \<rat>"
+ then obtain m n where
+ n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
+ and gcd: "gcd (m, n) = 1" ..
+ from n and sqrt_rat have "real m = \<bar>sqrt (real p)\<bar> * real n" by simp
+ then have "real (m\<twosuperior>) = (sqrt (real p))\<twosuperior> * real (n\<twosuperior>)"
+ by (auto simp add: power2_eq_square)
+ also have "(sqrt (real p))\<twosuperior> = real p" by simp
+ also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
+ finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
+ then have "p dvd m\<twosuperior>" ..
+ 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\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
+ with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
+ then have "p dvd n\<twosuperior>" ..
+ 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
+ then have "p \<le> 1" by (simp add: dvd_imp_le)
+ with p show False by simp
+qed
+
+end
diff --git a/isar/Root2_Tactic.thy b/isar/Root2_Tactic.thy
new file mode 100644
index 00000000..65f2a944
--- /dev/null
+++ b/isar/Root2_Tactic.thy
@@ -0,0 +1,81 @@
+(* Example proof by Larry Paulson; see http://www.cs.kun.nl/~freek/comparison/
+ Taken from Isabelle2004 distribution. *)
+
+
+(* Title: HOL/Hyperreal/ex/Sqrt_Script.thy
+ ID: Id: Sqrt_Script.thy,v 1.3 2003/12/10 14:59:35 paulson Exp
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Copyright 2001 University of Cambridge
+*)
+
+header {* Square roots of primes are irrational (script version) *}
+
+theory Sqrt_Script = Primes + Complex_Main:
+
+text {*
+ \medskip Contrast this linear Isabelle/Isar script with Markus
+ Wenzel's more mathematical version.
+*}
+
+subsection {* Preliminaries *}
+
+lemma prime_nonzero: "p \<in> prime \<Longrightarrow> p \<noteq> 0"
+ by (force simp add: prime_def)
+
+lemma prime_dvd_other_side:
+ "n * n = p * (k * k) \<Longrightarrow> p \<in> prime \<Longrightarrow> 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 \<in> prime \<Longrightarrow>
+ 0 < k \<Longrightarrow> k * k = p * (j * j) \<Longrightarrow> k < p * j \<and> 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 \<Longrightarrow> k * k = p * (j * j)"
+ by (simp add: mult_ac)
+
+lemma prime_not_square:
+ "p \<in> prime \<Longrightarrow> (\<And>k. 0 < k \<Longrightarrow> m * m \<noteq> p * (k * k))"
+ apply (induct 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" ("\<rat>")
+ "\<rat> \<equiv> {x. \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = 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 \<in> prime \<Longrightarrow> x * x = real p \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<notin> \<rat>"
+ 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: divide_eq_eq prime_not_square real_of_nat_mult [symmetric])
+ done
+
+lemmas two_sqrt_irrational =
+ prime_sqrt_irrational [OF two_is_prime]
+
+end