Theory Sqrt_Script
section ‹Square roots of primes are irrational (script version)›
text ‹
Contrast this linear Isabelle/Isar script with the more mathematical version
in 🗏‹~~/src/HOL/Examples/Sqrt.thy› by Makarius Wenzel.
›
theory Sqrt_Script
imports Complex_Main "HOL-Computational_Algebra.Primes"
begin
subsection ‹Preliminaries›
lemma prime_nonzero: "prime (p::nat) ⟹ p ≠ 0"
by (force simp add: prime_nat_iff)
lemma prime_dvd_other_side:
"(n::nat) * n = p * (k * k) ⟹ prime p ⟹ p dvd n"
apply (subgoal_tac "p dvd n * n", blast dest: prime_dvd_mult_nat)
apply auto
done
lemma reduction: "prime (p::nat) ⟹
0 < k ⟹ k * k = p * (j * j) ⟹ k < p * j ∧ 0 < j"
apply (rule ccontr)
apply (simp add: linorder_not_less)
apply (erule disjE)
apply (frule mult_le_mono, assumption)
apply auto
apply (force simp add: prime_nat_iff)
done
lemma rearrange: "(j::nat) * (p * j) = k * k ⟹ k * k = p * (j * j)"
by (simp add: ac_simps)
lemma prime_not_square:
"prime (p::nat) ⟹ (⋀k. 0 < k ⟹ m * m ≠ p * (k * k))"
apply (induct m rule: nat_less_induct)
apply clarify
apply (frule prime_dvd_other_side, assumption)
apply (erule dvdE)
apply (simp add: nat_mult_eq_cancel_disj prime_nonzero)
apply (blast dest: rearrange reduction)
done
subsection ‹Main theorem›
text ‹
The square root of any prime number (including ‹2›) is
irrational.
›
theorem prime_sqrt_irrational:
"prime (p::nat) ⟹ x * x = real p ⟹ 0 ≤ x ⟹ x ∉ ℚ"
apply (rule notI)
apply (erule Rats_abs_nat_div_natE)
apply (simp del: of_nat_mult
add: abs_if divide_eq_eq prime_not_square of_nat_mult [symmetric])
done
lemmas two_sqrt_irrational =
prime_sqrt_irrational [OF two_is_prime_nat]
end