src/HOL/ex/Sqrt.thy
changeset 73810 1c5dcba6925f
parent 73809 ce9529a616fd
equal deleted inserted replaced
73809:ce9529a616fd 73810:1c5dcba6925f
     1 (*  Title:      HOL/ex/Sqrt.thy
     1 (*  Title:      HOL/ex/Sqrt.thy
     2     Author:     Markus Wenzel, Tobias Nipkow, TU Muenchen
     2     Author:     Makarius
       
     3     Author:     Tobias Nipkow, TU Muenchen
     3 *)
     4 *)
     4 
     5 
     5 section \<open>Square roots of primes are irrational\<close>
     6 section \<open>Square roots of primes are irrational\<close>
     6 
     7 
     7 theory Sqrt
     8 theory Sqrt
     8 imports Complex_Main "HOL-Computational_Algebra.Primes"
     9   imports Complex_Main "HOL-Computational_Algebra.Primes"
     9 begin
    10 begin
    10 
    11 
    11 text \<open>The square root of any prime number (including 2) is irrational.\<close>
    12 text \<open>
       
    13   The square root of any prime number (including 2) is irrational.
       
    14 \<close>
    12 
    15 
    13 theorem sqrt_prime_irrational:
    16 theorem sqrt_prime_irrational:
    14   assumes "prime (p::nat)"
    17   fixes p :: nat
       
    18   assumes "prime p"
    15   shows "sqrt p \<notin> \<rat>"
    19   shows "sqrt p \<notin> \<rat>"
    16 proof
    20 proof
    17   from \<open>prime p\<close> have p: "1 < p" by (simp add: prime_nat_iff)
    21   from \<open>prime p\<close> have p: "p > 1" by (rule prime_gt_1_nat)
    18   assume "sqrt p \<in> \<rat>"
    22   assume "sqrt p \<in> \<rat>"
    19   then obtain m n :: nat where
    23   then obtain m n :: nat
    20       n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt p\<bar> = m / n"
    24     where n: "n \<noteq> 0"
    21     and "coprime m n" by (rule Rats_abs_nat_div_natE)
    25       and sqrt_rat: "\<bar>sqrt p\<bar> = m / n"
       
    26       and "coprime m n" by (rule Rats_abs_nat_div_natE)
    22   have eq: "m\<^sup>2 = p * n\<^sup>2"
    27   have eq: "m\<^sup>2 = p * n\<^sup>2"
    23   proof -
    28   proof -
    24     from n and sqrt_rat have "m = \<bar>sqrt p\<bar> * n" by simp
    29     from n and sqrt_rat have "m = \<bar>sqrt p\<bar> * n" by simp
    25     then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2"
    30     then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2" by (simp add: power_mult_distrib)
    26       by (auto simp add: power2_eq_square)
       
    27     also have "(sqrt p)\<^sup>2 = p" by simp
    31     also have "(sqrt p)\<^sup>2 = p" by simp
    28     also have "\<dots> * n\<^sup>2 = p * n\<^sup>2" by simp
    32     also have "\<dots> * n\<^sup>2 = p * n\<^sup>2" by simp
    29     finally show ?thesis using of_nat_eq_iff by blast
    33     finally show ?thesis by linarith
    30   qed
    34   qed
    31   have "p dvd m \<and> p dvd n"
    35   have "p dvd m \<and> p dvd n"
    32   proof
    36   proof
    33     from eq have "p dvd m\<^sup>2" ..
    37     from eq have "p dvd m\<^sup>2" ..
    34     with \<open>prime p\<close> show "p dvd m" by (rule prime_dvd_power_nat)
    38     with \<open>prime p\<close> show "p dvd m" by (rule prime_dvd_power)
    35     then obtain k where "m = p * k" ..
    39     then obtain k where "m = p * k" ..
    36     with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square ac_simps)
    40     with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by algebra
    37     with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
    41     with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
    38     then have "p dvd n\<^sup>2" ..
    42     then have "p dvd n\<^sup>2" ..
    39     with \<open>prime p\<close> show "p dvd n" by (rule prime_dvd_power_nat)
    43     with \<open>prime p\<close> show "p dvd n" by (rule prime_dvd_power)
    40   qed
    44   qed
    41   then have "p dvd gcd m n" by simp
    45   then have "p dvd gcd m n" by simp
    42   with \<open>coprime m n\<close> have "p = 1" by simp
    46   with \<open>coprime m n\<close> have "p = 1" by simp
    43   with p show False by simp
    47   with p show False by simp
    44 qed
    48 qed
    45 
    49 
    46 corollary sqrt_2_not_rat: "sqrt 2 \<notin> \<rat>"
    50 corollary sqrt_2_not_rat: "sqrt 2 \<notin> \<rat>"
    47   using sqrt_prime_irrational[of 2] by simp
    51   using sqrt_prime_irrational [of 2] by simp
    48 
       
    49 
       
    50 subsection \<open>Variations\<close>
       
    51 
    52 
    52 text \<open>
    53 text \<open>
    53   Here is an alternative version of the main proof, using mostly
    54   Here is an alternative version of the main proof, using mostly linear
    54   linear forward-reasoning.  While this results in less top-down
    55   forward-reasoning. While this results in less top-down structure, it is
    55   structure, it is probably closer to proofs seen in mathematics.
    56   probably closer to proofs seen in mathematics.
    56 \<close>
    57 \<close>
    57 
    58 
    58 theorem
    59 theorem
    59   assumes "prime (p::nat)"
    60   fixes p :: nat
       
    61   assumes "prime p"
    60   shows "sqrt p \<notin> \<rat>"
    62   shows "sqrt p \<notin> \<rat>"
    61 proof
    63 proof
    62   from \<open>prime p\<close> have p: "1 < p" by (simp add: prime_nat_iff)
    64   from \<open>prime p\<close> have p: "p > 1" by (rule prime_gt_1_nat)
    63   assume "sqrt p \<in> \<rat>"
    65   assume "sqrt p \<in> \<rat>"
    64   then obtain m n :: nat where
    66   then obtain m n :: nat
    65       n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt p\<bar> = m / n"
    67     where n: "n \<noteq> 0"
    66     and "coprime m n" by (rule Rats_abs_nat_div_natE)
    68       and sqrt_rat: "\<bar>sqrt p\<bar> = m / n"
       
    69       and "coprime m n" by (rule Rats_abs_nat_div_natE)
    67   from n and sqrt_rat have "m = \<bar>sqrt p\<bar> * n" by simp
    70   from n and sqrt_rat have "m = \<bar>sqrt p\<bar> * n" by simp
    68   then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2"
    71   then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2" by (auto simp add: power2_eq_square)
    69     by (auto simp add: power2_eq_square)
       
    70   also have "(sqrt p)\<^sup>2 = p" by simp
    72   also have "(sqrt p)\<^sup>2 = p" by simp
    71   also have "\<dots> * n\<^sup>2 = p * n\<^sup>2" by simp
    73   also have "\<dots> * n\<^sup>2 = p * n\<^sup>2" by simp
    72   finally have eq: "m\<^sup>2 = p * n\<^sup>2" using of_nat_eq_iff by blast
    74   finally have eq: "m\<^sup>2 = p * n\<^sup>2" by linarith
    73   then have "p dvd m\<^sup>2" ..
    75   then have "p dvd m\<^sup>2" ..
    74   with \<open>prime p\<close> have dvd_m: "p dvd m" by (rule prime_dvd_power_nat)
    76   with \<open>prime p\<close> have dvd_m: "p dvd m" by (rule prime_dvd_power)
    75   then obtain k where "m = p * k" ..
    77   then obtain k where "m = p * k" ..
    76   with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square ac_simps)
    78   with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by algebra
    77   with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
    79   with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
    78   then have "p dvd n\<^sup>2" ..
    80   then have "p dvd n\<^sup>2" ..
    79   with \<open>prime p\<close> have "p dvd n" by (rule prime_dvd_power_nat)
    81   with \<open>prime p\<close> have "p dvd n" by (rule prime_dvd_power)
    80   with dvd_m have "p dvd gcd m n" by (rule gcd_greatest)
    82   with dvd_m have "p dvd gcd m n" by (rule gcd_greatest)
    81   with \<open>coprime m n\<close> have "p = 1" by simp
    83   with \<open>coprime m n\<close> have "p = 1" by simp
    82   with p show False by simp
    84   with p show False by simp
    83 qed
    85 qed
    84 
    86 
    85 
    87 
    86 text \<open>Another old chestnut, which is a consequence of the irrationality of \<^term>\<open>sqrt 2\<close>.\<close>
    88 text \<open>
       
    89   Another old chestnut, which is a consequence of the irrationality of
       
    90   \<^term>\<open>sqrt 2\<close>.
       
    91 \<close>
    87 
    92 
    88 lemma "\<exists>a b::real. a \<notin> \<rat> \<and> b \<notin> \<rat> \<and> a powr b \<in> \<rat>" (is "\<exists>a b. ?P a b")
    93 lemma "\<exists>a b::real. a \<notin> \<rat> \<and> b \<notin> \<rat> \<and> a powr b \<in> \<rat>" (is "\<exists>a b. ?P a b")
    89 proof (cases "sqrt 2 powr sqrt 2 \<in> \<rat>")
    94 proof (cases "sqrt 2 powr sqrt 2 \<in> \<rat>")
    90   case True
    95   case True
    91   with sqrt_2_not_rat have "?P (sqrt 2) (sqrt 2)" by simp
    96   with sqrt_2_not_rat have "?P (sqrt 2) (sqrt 2)" by simp