src/HOL/Complex/ex/Sqrt.thy
author wenzelm
Wed, 09 Jun 2004 18:52:42 +0200
changeset 14898 a25550451b51
parent 14353 79f9fbef9106
child 14981 e73f8140af78
permissions -rw-r--r--
Url.File;
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
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
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13957
diff changeset
     9
theory Sqrt = Primes + Complex_Main:
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    10
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
subsection {* Preliminaries *}
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
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    14
  The set of rational numbers, including the key representation
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
  theorem.
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    18
constdefs
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    19
  rationals :: "real set"    ("\<rat>")
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
  "\<rat> \<equiv> {x. \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real (m::nat) / real (n::nat)}"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
theorem rationals_rep: "x \<in> \<rat> \<Longrightarrow>
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
  \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real m / real n \<and> gcd (m, n) = 1"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    24
proof -
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
  assume "x \<in> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    26
  then obtain m n :: nat where
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
      n: "n \<noteq> 0" and x_rat: "\<bar>x\<bar> = real m / real n"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
    by (unfold rationals_def) blast
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
  let ?gcd = "gcd (m, n)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
  from n have gcd: "?gcd \<noteq> 0" by (simp add: gcd_zero)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    31
  let ?k = "m div ?gcd"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    32
  let ?l = "n div ?gcd"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
  let ?gcd' = "gcd (?k, ?l)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
  have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
    by (rule dvd_mult_div_cancel)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
  have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n"
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
  from n and gcd_l have "?l \<noteq> 0"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
    by (auto iff del: neq0_conv)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
  moreover
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
  have "\<bar>x\<bar> = real ?k / real ?l"
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 gcd have "real ?k / real ?l =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    45
        real (?gcd * ?k) / real (?gcd * ?l)"
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14051
diff changeset
    46
      by (simp add: mult_divide_cancel_left)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    47
    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
    48
    also from x_rat have "\<dots> = \<bar>x\<bar>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
    finally show ?thesis ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
  moreover
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    52
  have "?gcd' = 1"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
  proof -
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
    have "?gcd * ?gcd' = gcd (?gcd * ?k, ?gcd * ?l)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
      by (rule gcd_mult_distrib2)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
    with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
    with gcd show ?thesis by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
  ultimately show ?thesis by blast
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
qed
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
lemma [elim?]: "r \<in> \<rat> \<Longrightarrow>
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    63
  (\<And>m n. n \<noteq> 0 \<Longrightarrow> \<bar>r\<bar> = real m / real n \<Longrightarrow> gcd (m, n) = 1 \<Longrightarrow> C)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    64
    \<Longrightarrow> C"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
  using rationals_rep by blast
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
subsection {* Main theorem *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    69
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    70
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
  The square root of any prime number (including @{text 2}) is
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    72
  irrational.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    73
*}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    74
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
theorem sqrt_prime_irrational: "p \<in> prime \<Longrightarrow> sqrt (real p) \<notin> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    76
proof
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    77
  assume p_prime: "p \<in> prime"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    78
  then have p: "1 < p" by (simp add: prime_def)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
  assume "sqrt (real p) \<in> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
  then obtain m n where
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    81
      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
    82
    and gcd: "gcd (m, n) = 1" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
  have eq: "m\<twosuperior> = p * n\<twosuperior>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
  proof -
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
    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
    86
    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
    87
      by (auto simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    88
    also have "(sqrt (real p))\<twosuperior> = real p" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
    also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
    finally show ?thesis ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    91
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
  have "p dvd m \<and> p dvd n"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
  proof
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
    from eq have "p dvd m\<twosuperior>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    95
    with p_prime show "p dvd m" by (rule prime_dvd_power_two)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    96
    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
    97
    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
    98
    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
    99
    then have "p dvd n\<twosuperior>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   100
    with p_prime show "p dvd n" by (rule prime_dvd_power_two)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   101
  qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   102
  then have "p dvd gcd (m, n)" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   103
  with gcd have "p dvd 1" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   104
  then have "p \<le> 1" by (simp add: dvd_imp_le)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   105
  with p show False by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   106
qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   107
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   108
corollary "sqrt (real (2::nat)) \<notin> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   109
  by (rule sqrt_prime_irrational) (rule two_is_prime)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   110
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
subsection {* Variations *}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   113
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   114
text {*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   115
  Here is an alternative version of the main proof, using mostly
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   116
  linear forward-reasoning.  While this results in less top-down
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   117
  structure, it is probably closer to proofs seen in mathematics.
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   118
*}
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   119
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   120
theorem "p \<in> prime \<Longrightarrow> sqrt (real p) \<notin> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   121
proof
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   122
  assume p_prime: "p \<in> prime"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   123
  then have p: "1 < p" by (simp add: prime_def)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   124
  assume "sqrt (real p) \<in> \<rat>"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   125
  then obtain m n where
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   126
      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
   127
    and gcd: "gcd (m, n) = 1" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   128
  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
   129
  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
   130
    by (auto simp add: power2_eq_square)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   131
  also have "(sqrt (real p))\<twosuperior> = real p" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   132
  also have "\<dots> * real (n\<twosuperior>) = real (p * n\<twosuperior>)" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   133
  finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   134
  then have "p dvd m\<twosuperior>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   135
  with p_prime have dvd_m: "p dvd m" by (rule prime_dvd_power_two)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   136
  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
   137
  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
   138
  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
   139
  then have "p dvd n\<twosuperior>" ..
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   140
  with p_prime have "p dvd n" by (rule prime_dvd_power_two)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   141
  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
   142
  with gcd have "p dvd 1" by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   143
  then have "p \<le> 1" by (simp add: dvd_imp_le)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   144
  with p show False by simp
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   145
qed
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   146
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   147
end