added HOL-Hyperreal-ex;
authorwenzelm
Wed, 06 Mar 2002 17:47:51 +0100
changeset 13029 84e4ba7fb033
parent 13028 81c87faed78b
child 13030 5765aa72afac
added HOL-Hyperreal-ex;
src/HOL/Hyperreal/ex/ROOT.ML
src/HOL/Hyperreal/ex/Sqrt.thy
src/HOL/Hyperreal/ex/Sqrt_Script.thy
src/HOL/Hyperreal/ex/document/root.tex
src/HOL/IsaMakefile
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/ex/ROOT.ML	Wed Mar 06 17:47:51 2002 +0100
@@ -0,0 +1,9 @@
+(*  Title:      HOL/Hyperreal/ex/ROOT.ML
+    ID:         $Id$
+
+Miscellaneous HOL-Hyperreal Examples.
+*)
+
+no_document use_thy "Primes";
+use_thy "Sqrt";
+use_thy "Sqrt_Script";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/ex/Sqrt.thy	Wed Mar 06 17:47:51 2002 +0100
@@ -0,0 +1,151 @@
+(*  Title:      HOL/Hyperreal/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 + Hyperreal:
+
+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: real_mult_div_cancel1)
+    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: power_two real_power_two)
+    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: power_two mult_ac)
+    with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power_two)
+    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
+
+text {*
+  Just for the record: we instantiate the main theorem for the
+  specific prime number @{text 2} (real mathematicians would never do
+  this formally :-).
+*}
+
+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: power_two real_power_two)
+  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: power_two mult_ac)
+  with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power_two)
+  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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/ex/Sqrt_Script.thy	Wed Mar 06 17:47:51 2002 +0100
@@ -0,0 +1,76 @@
+(*  Title:      HOL/Hyperreal/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 (script version) *}
+
+theory Sqrt_Script = Primes + Hyperreal:
+
+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: real_divide_eq_eq prime_not_square
+    real_of_nat_mult [symmetric])
+  done
+
+lemmas two_sqrt_irrational = prime_sqrt_irrational [OF two_is_prime]
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/ex/document/root.tex	Wed Mar 06 17:47:51 2002 +0100
@@ -0,0 +1,20 @@
+
+\documentclass[11pt,a4paper]{article}
+\usepackage[latin1]{inputenc}
+\usepackage{isabelle,isabellesym}
+\usepackage{pdfsetup}
+
+\urlstyle{rm}
+\isabellestyle{it}
+
+\begin{document}
+
+\title{Miscellaneous HOL-Hyperreal Examples}
+\maketitle
+
+\tableofcontents
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\end{document}
--- a/src/HOL/IsaMakefile	Wed Mar 06 16:18:45 2002 +0100
+++ b/src/HOL/IsaMakefile	Wed Mar 06 17:47:51 2002 +0100
@@ -22,6 +22,7 @@
       HOL-Real-ex \
   HOL-Hoare \
   HOL-HoareParallel \
+  HOL-Hyperreal-ex \
   HOL-IMP \
   HOL-IMPP \
   HOL-IOA \
@@ -123,7 +124,33 @@
 	@cd Real; $(ISATOOL) usedir -b $(OUT)/HOL HOL-Real
 
 
-## HOL-Real-Hyperreal
+## HOL-Real-ex
+
+HOL-Real-ex: HOL-Real $(LOG)/HOL-Real-ex.gz
+
+$(LOG)/HOL-Real-ex.gz: $(OUT)/HOL-Real Real/ex/ROOT.ML \
+  Real/ex/BinEx.thy Real/ex/document/root.tex
+	@cd Real; $(ISATOOL) usedir $(OUT)/HOL-Real ex
+
+
+## HOL-Real-HahnBanach
+
+HOL-Real-HahnBanach: HOL-Real $(LOG)/HOL-Real-HahnBanach.gz
+
+$(LOG)/HOL-Real-HahnBanach.gz: $(OUT)/HOL-Real Real/HahnBanach/Aux.thy \
+  Real/HahnBanach/Bounds.thy Real/HahnBanach/FunctionNorm.thy \
+  Real/HahnBanach/FunctionOrder.thy Real/HahnBanach/HahnBanach.thy \
+  Real/HahnBanach/HahnBanachExtLemmas.thy	\
+  Real/HahnBanach/HahnBanachSupLemmas.thy	\
+  Real/HahnBanach/Linearform.thy Real/HahnBanach/NormedSpace.thy \
+  Real/HahnBanach/README.html Real/HahnBanach/ROOT.ML \
+  Real/HahnBanach/Subspace.thy Real/HahnBanach/VectorSpace.thy \
+  Real/HahnBanach/ZornLemma.thy Real/HahnBanach/document/root.bib \
+  Real/HahnBanach/document/root.tex
+	@cd Real; $(ISATOOL) usedir -g true $(OUT)/HOL-Real HahnBanach
+
+
+## HOL-Hyperreal
 
 HOL-Hyperreal: HOL-Real $(OUT)/HOL-Hyperreal
 
@@ -151,31 +178,14 @@
 	@cd Hyperreal; $(ISATOOL) usedir -b $(OUT)/HOL-Real HOL-Hyperreal
 
 
-## HOL-Real-ex
-
-HOL-Real-ex: HOL-Real $(LOG)/HOL-Real-ex.gz
+## HOL-Hyperreal-ex
 
-$(LOG)/HOL-Real-ex.gz: $(OUT)/HOL-Real Real/ex/ROOT.ML \
-  Real/ex/BinEx.thy Real/ex/document/root.tex Real/ex/Sqrt.thy \
-  Real/ex/Sqrt_Script.thy
-	@cd Real; $(ISATOOL) usedir $(OUT)/HOL-Real ex
-
-
-## HOL-Real-HahnBanach
+HOL-Hyperreal-ex: HOL-Hyperreal $(LOG)/HOL-Hyperreal-ex.gz
 
-HOL-Real-HahnBanach: HOL-Real $(LOG)/HOL-Real-HahnBanach.gz
-
-$(LOG)/HOL-Real-HahnBanach.gz: $(OUT)/HOL-Real Real/HahnBanach/Aux.thy \
-  Real/HahnBanach/Bounds.thy Real/HahnBanach/FunctionNorm.thy \
-  Real/HahnBanach/FunctionOrder.thy Real/HahnBanach/HahnBanach.thy \
-  Real/HahnBanach/HahnBanachExtLemmas.thy	\
-  Real/HahnBanach/HahnBanachSupLemmas.thy	\
-  Real/HahnBanach/Linearform.thy Real/HahnBanach/NormedSpace.thy \
-  Real/HahnBanach/README.html Real/HahnBanach/ROOT.ML \
-  Real/HahnBanach/Subspace.thy Real/HahnBanach/VectorSpace.thy \
-  Real/HahnBanach/ZornLemma.thy Real/HahnBanach/document/root.bib \
-  Real/HahnBanach/document/root.tex
-	@cd Real; $(ISATOOL) usedir -g true $(OUT)/HOL-Real HahnBanach
+$(LOG)/HOL-Hyperreal-ex.gz: $(OUT)/HOL-Hyperreal Hyperreal/ex/ROOT.ML \
+  Hyperreal/ex/document/root.tex Hyperreal/ex/Sqrt.thy \
+  Hyperreal/ex/Sqrt_Script.thy
+	@cd Hyperreal; $(ISATOOL) usedir $(OUT)/HOL-Hyperreal ex
 
 
 ## HOL-Library
@@ -301,7 +311,7 @@
   HoareParallel/RG_Hoare.thy HoareParallel/RG_Syntax.thy	   \
   HoareParallel/RG_Tran.thy HoareParallel/ROOT.ML		   \
   HoareParallel/document/root.tex
-	@$(ISATOOL) usedir $(OUT)/HOL HoareParallel
+	@$(ISATOOL) usedir -g true $(OUT)/HOL HoareParallel
 
 
 ## HOL-Lex
@@ -652,6 +662,7 @@
 		$(LOG)/HOL-MicroJava.gz $(LOG)/HOL-NanoJava.gz \
 		$(LOG)/HOL-IOA.gz $(LOG)/HOL-AxClasses \
 		$(LOG)/HOL-Lattice $(LOG)/HOL-Real-ex.gz \
+		$(LOG)/HOL-Hyperreal-ex.gz \
 		$(LOG)/HOL-Real-HahnBanach.gz $(LOG)/TLA-Inc.gz \
 		$(LOG)/TLA-Buffer.gz $(LOG)/TLA-Memory.gz \
 		$(LOG)/HOL-Library.gz $(LOG)/HOL-Unix.gz