src/HOL/Complex/ex/Sqrt.thy
author kleing
Sun, 19 Feb 2006 13:21:32 +0100
changeset 19106 6e6b5b1fdc06
parent 19086 1b3780be6cc2
child 21404 eb85850d3eb7
permissions -rw-r--r--
* added Library/ASeries (sum of arithmetic series with instantiation to nat and int) * added Complex/ex/ASeries_Complex (instantiation of the above for reals) * added Complex/ex/HarmonicSeries (should really be in something like Complex/Library) (these are contributions by Benjamin Porter, numbers 68 and 34 of http://www.cs.ru.nl/~freek/100/)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Hyperreal/ex/Sqrt.thy
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     2
    ID:         $Id$
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
14981
e73f8140af78 Merged in license change from Isabelle2004
kleing
parents: 14353
diff changeset
     4
13957
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     7
header {*  Square roots of primes are irrational *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     8
15149
c5c4884634b7 new import syntax
nipkow
parents: 14981
diff changeset
     9
theory Sqrt
c5c4884634b7 new import syntax
nipkow
parents: 14981
diff changeset
    10
imports Primes Complex_Main
c5c4884634b7 new import syntax
nipkow
parents: 14981
diff changeset
    11
begin
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    12
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
subsection {* Preliminaries *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    14
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
  The set of rational numbers, including the key representation
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    17
  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
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    20
definition
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    21
  rationals  ("\<rat>")
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    22
  "\<rat> = {x. \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real (m::nat) / real (n::nat)}"
13957
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 rationals_rep [elim?]:
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    25
  assumes "x \<in> \<rat>"
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    26
  obtains m n where "n \<noteq> 0" and "\<bar>x\<bar> = real m / real n" and "gcd (m, n) = 1"
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 `x \<in> \<rat>` obtain m n :: nat where
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
      n: "n \<noteq> 0" and x_rat: "\<bar>x\<bar> = real m / real n"
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    30
    unfolding rationals_def by blast
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    31
  let ?gcd = "gcd (m, n)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    32
  from n have gcd: "?gcd \<noteq> 0" by (simp add: gcd_zero)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
  let ?k = "m div ?gcd"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
  let ?l = "n div ?gcd"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
  let ?gcd' = "gcd (?k, ?l)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
  have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
    by (rule dvd_mult_div_cancel)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
  have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
    by (rule dvd_mult_div_cancel)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
  from n and gcd_l have "?l \<noteq> 0"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
    by (auto iff del: neq0_conv)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
  moreover
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
  have "\<bar>x\<bar> = real ?k / real ?l"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    45
  proof -
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
    from gcd have "real ?k / real ?l =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    47
        real (?gcd * ?k) / real (?gcd * ?l)"
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14051
diff changeset
    48
      by (simp add: mult_divide_cancel_left)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
    also from gcd_k and gcd_l have "\<dots> = real m / real n" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
    also from x_rat have "\<dots> = \<bar>x\<bar>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
    finally show ?thesis ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    52
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
  moreover
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
  have "?gcd' = 1"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
  proof -
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
    have "?gcd * ?gcd' = gcd (?gcd * ?k, ?gcd * ?l)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
      by (rule gcd_mult_distrib2)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
    with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
    with gcd show ?thesis by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
  qed
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    61
  ultimately show ?thesis ..
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    62
qed
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
subsection {* Main theorem *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    66
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    67
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
  The square root of any prime number (including @{text 2}) is
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    69
  irrational.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    70
*}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    72
theorem sqrt_prime_irrational:
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    73
  assumes "prime p"
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    74
  shows "sqrt (real p) \<notin> \<rat>"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
proof
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    76
  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
    77
  assume "sqrt (real p) \<in> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    78
  then obtain m n where
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
      n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
    and gcd: "gcd (m, n) = 1" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    81
  have eq: "m\<twosuperior> = p * n\<twosuperior>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    82
  proof -
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
    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
    84
    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
    85
      by (auto simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    86
    also have "(sqrt (real p))\<twosuperior> = real p" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    87
    also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    88
    finally show ?thesis ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
  have "p dvd m \<and> p dvd n"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    91
  proof
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
    from eq have "p dvd m\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    93
    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
    94
    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
    95
    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
    96
    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
    97
    then have "p dvd n\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
    98
    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
    99
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   100
  then have "p dvd gcd (m, n)" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   101
  with gcd have "p dvd 1" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   102
  then have "p \<le> 1" by (simp add: dvd_imp_le)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   103
  with p show False by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   104
qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   105
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   106
corollary "sqrt (real (2::nat)) \<notin> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   107
  by (rule sqrt_prime_irrational) (rule two_is_prime)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   108
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   109
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   110
subsection {* Variations *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   111
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   112
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   113
  Here is an alternative version of the main proof, using mostly
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   114
  linear forward-reasoning.  While this results in less top-down
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   115
  structure, it is probably closer to proofs seen in mathematics.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   116
*}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   117
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
   118
theorem
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
   119
  assumes "prime p"
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
   120
  shows "sqrt (real p) \<notin> \<rat>"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   121
proof
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
   122
  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
   123
  assume "sqrt (real p) \<in> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   124
  then obtain m n where
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   125
      n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt (real p)\<bar> = real m / real n"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   126
    and gcd: "gcd (m, n) = 1" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   127
  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
   128
  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
   129
    by (auto simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   130
  also have "(sqrt (real p))\<twosuperior> = real p" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   131
  also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   132
  finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   133
  then have "p dvd m\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
   134
  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
   135
  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
   136
  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
   137
  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
   138
  then have "p dvd n\<twosuperior>" ..
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16663
diff changeset
   139
  with `prime p` have "p dvd n" by (rule prime_dvd_power_two)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   140
  with dvd_m have "p dvd gcd (m, n)" by (rule gcd_greatest)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   141
  with gcd have "p dvd 1" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   142
  then have "p \<le> 1" by (simp add: dvd_imp_le)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   143
  with p show False by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   144
qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   145
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   146
end