src/HOL/ex/Sqrt.thy
author blanchet
Wed, 04 Mar 2009 11:05:29 +0100
changeset 30242 aea5d7fa7ef5
parent 28952 15a4b2cf8c34
child 30411 9c9b6511ad1b
permissions -rw-r--r--
Merge.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28001
diff changeset
     1
(*  Title:      HOL/ex/Sqrt.thy
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
14981
e73f8140af78 Merged in license change from Isabelle2004
kleing
parents: 14353
diff changeset
     3
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     6
header {*  Square roots of primes are irrational *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     7
15149
c5c4884634b7 new import syntax
nipkow
parents: 14981
diff changeset
     8
theory Sqrt
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28001
diff changeset
     9
imports Complex_Main Primes
15149
c5c4884634b7 new import syntax
nipkow
parents: 14981
diff changeset
    10
begin
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27556
diff changeset
    12
text {* The definition and the key representation theorem for the set of
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27556
diff changeset
    13
rational numbers has been moved to a core theory.  *}
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    14
28001
4642317e0deb Defined rationals (Rats) globally in Rational.
nipkow
parents: 27556
diff changeset
    15
declare Rats_abs_nat_div_natE[elim?]
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    17
subsection {* Main theorem *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    18
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    19
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
  The square root of any prime number (including @{text 2}) is
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
  irrational.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
*}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    24
theorem sqrt_prime_irrational:
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    25
  assumes "prime p"
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    26
  shows "sqrt (real p) \<notin> \<rat>"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
proof
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    28
  from `prime p` have p: "1 < p" by (simp add: prime_def)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
  assume "sqrt (real p) \<in> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
  then obtain m n where
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    31
      n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 23413
diff changeset
    32
    and gcd: "gcd m n = 1" ..
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
  have eq: "m\<twosuperior> = p * n\<twosuperior>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
  proof -
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
    from n and sqrt_rat have "real m = \<bar>sqrt (real p)\<bar> * real n" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
    then have "real (m\<twosuperior>) = (sqrt (real p))\<twosuperior> * real (n\<twosuperior>)"
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14305
diff changeset
    37
      by (auto simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
    also have "(sqrt (real p))\<twosuperior> = real p" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
    also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
    finally show ?thesis ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
  have "p dvd m \<and> p dvd n"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
  proof
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
    from eq have "p dvd m\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    45
    with `prime p` show "p dvd m" by (rule prime_dvd_power_two)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
    then obtain k where "m = p * k" ..
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14305
diff changeset
    47
    with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14305
diff changeset
    48
    with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
    then have "p dvd n\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    50
    with `prime p` show "p dvd n" by (rule prime_dvd_power_two)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
  qed
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 23413
diff changeset
    52
  then have "p dvd gcd m n" ..
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
  with gcd have "p dvd 1" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
  then have "p \<le> 1" by (simp add: dvd_imp_le)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
  with p show False by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
corollary "sqrt (real (2::nat)) \<notin> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
  by (rule sqrt_prime_irrational) (rule two_is_prime)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    61
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    62
subsection {* Variations *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    63
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    64
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
  Here is an alternative version of the main proof, using mostly
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    66
  linear forward-reasoning.  While this results in less top-down
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    67
  structure, it is probably closer to proofs seen in mathematics.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
*}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    69
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    70
theorem
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    71
  assumes "prime p"
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    72
  shows "sqrt (real p) \<notin> \<rat>"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    73
proof
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    74
  from `prime p` have p: "1 < p" by (simp add: prime_def)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
  assume "sqrt (real p) \<in> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    76
  then obtain m n where
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    77
      n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 23413
diff changeset
    78
    and gcd: "gcd m n = 1" ..
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
  from n and sqrt_rat have "real m = \<bar>sqrt (real p)\<bar> * real n" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
  then have "real (m\<twosuperior>) = (sqrt (real p))\<twosuperior> * real (n\<twosuperior>)"
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14305
diff changeset
    81
    by (auto simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    82
  also have "(sqrt (real p))\<twosuperior> = real p" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
  also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
  finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
  then have "p dvd m\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    86
  with `prime p` have dvd_m: "p dvd m" by (rule prime_dvd_power_two)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    87
  then obtain k where "m = p * k" ..
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14305
diff changeset
    88
  with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14305
diff changeset
    89
  with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
  then have "p dvd n\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    91
  with `prime p` have "p dvd n" by (rule prime_dvd_power_two)
27556
292098f2efdf unified curried gcd, lcm, zgcd, zlcm
haftmann
parents: 23413
diff changeset
    92
  with dvd_m have "p dvd gcd m n" by (rule gcd_greatest)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
  with gcd have "p dvd 1" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
  then have "p \<le> 1" by (simp add: dvd_imp_le)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    95
  with p show False by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    96
qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    97
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    98
end