src/HOL/Real/ex/Sqrt_Script.thy
author kleing
Thu, 17 Jan 2002 15:06:36 +0100
changeset 12791 ccc0f45ad2c4
parent 12087 b38cfbabfda4
permissions -rw-r--r--
registered directly executable version with the code generator
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Real/ex/Sqrt_Script.thy
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     2
    ID:         $Id$
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     5
*)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     6
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     7
header {*  Square roots of primes are irrational *}
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
     8
12077
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
     9
theory Sqrt_Script = Primes + Real:
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
    10
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    11
text {*
12077
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
    12
  \medskip Contrast this linear Isar script with Markus Wenzel's more
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
    13
  mathematical version.
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    14
*}
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    15
12087
wenzelm
parents: 12077
diff changeset
    16
subsection {* Preliminaries *}
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    17
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    18
lemma prime_nonzero:  "p \<in> prime \<Longrightarrow> p\<noteq>0"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    19
by (force simp add: prime_def)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    20
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    21
lemma prime_dvd_other_side: "\<lbrakk>n*n = p*(k*k); p \<in> prime\<rbrakk> \<Longrightarrow> p dvd n"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    22
apply (subgoal_tac "p dvd n*n", blast dest: prime_dvd_mult)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    23
apply (rule_tac j="k*k" in dvd_mult_left, simp)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    24
done
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    25
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    26
lemma reduction: "\<lbrakk>p \<in> prime; 0 < k; k*k = p*(j*j)\<rbrakk> \<Longrightarrow> k < p*j & 0 < j"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    27
apply (rule ccontr)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    28
apply (simp add: linorder_not_less)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    29
apply (erule disjE)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    30
 apply (frule mult_le_mono, assumption)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    31
 apply auto
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    32
apply (force simp add: prime_def)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    33
done
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    34
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    35
lemma rearrange: "(j::nat) * (p*j) = k*k \<Longrightarrow> k*k = p*(j*j)"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    36
by (simp add: mult_ac)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    37
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    38
lemma prime_not_square [rule_format]:
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    39
     "p \<in> prime \<Longrightarrow> \<forall>k. 0<k \<longrightarrow> m*m \<noteq> p*(k*k)"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    40
apply (induct_tac m rule: nat_less_induct)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    41
apply clarify 
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    42
apply (frule prime_dvd_other_side, assumption)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    43
apply (erule dvdE)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    44
apply (simp add: nat_mult_eq_cancel_disj prime_nonzero)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    45
apply (blast dest: rearrange reduction)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    46
done
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    47
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    48
12087
wenzelm
parents: 12077
diff changeset
    49
subsection {* The set of rational numbers *}
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    50
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    51
constdefs
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    52
  rationals :: "real set"    ("\<rat>")
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    53
  "\<rat> \<equiv> {x. \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real (m::nat) / real (n::nat)}"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    54
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    55
12087
wenzelm
parents: 12077
diff changeset
    56
subsection {* Main theorem *}
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    57
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    58
text {*
12077
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
    59
  The square root of any prime number (including @{text 2}) is
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
    60
  irrational.
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    61
*}
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    62
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    63
theorem prime_sqrt_irrational: "\<lbrakk>p \<in> prime; x*x = real p; 0 \<le> x\<rbrakk> \<Longrightarrow> x \<notin> \<rat>"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    64
apply (simp add: rationals_def real_abs_def)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    65
apply clarify
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    66
apply (erule_tac P="real m / real n * ?x = ?y" in rev_mp) 
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    67
apply (simp del: real_of_nat_mult
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    68
	    add: real_divide_eq_eq prime_not_square
12077
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
    69
                 real_of_nat_mult [symmetric])
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    70
done
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    71
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    72
lemma two_is_prime: "2 \<in> prime"
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    73
apply (auto simp add: prime_def)
12077
d46a32262bac activate dead code, make document work;
wenzelm
parents: 12051
diff changeset
    74
apply (case_tac m)
12051
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    75
apply (auto dest!: dvd_imp_le)
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    76
apply arith
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    77
done
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    78
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    79
lemmas two_sqrt_irrational = prime_sqrt_irrational [OF two_is_prime]
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    80
650f854b7310 new Sqrt example
paulson
parents:
diff changeset
    81
end