Computing Square Roots using the Babylonian

Computing Square Roots using the Babylonian
Method∗
René Thiemann
February 16, 2013
Abstract
We implement the Babylonian method [1] to compute square roots
of numbers. We provide precise algorithms for naturals, integers and
rationals, and offer an approximation algorithm for linear ordered
fields.
Contents
1 Introduction
2
2 The Babylonian method
2
3 The Babylonian method using integer division
2
4 Square roots for the naturals
9
5 Square roots for the rationals
10
6 Approximating square roots
12
7 Some tests
16
theory Sqrt-Babylonian
imports Rat
begin
∗
This research is supported by FWF (Austrian Science Fund) project P22767-N13.
1
1
Introduction
This theory provides executable algorithms for computing square-roots of
numbers which are all based on the Babylonian method (which is also known
as Heron’s method or Newton’s method).
For integers / naturals / rationals precise algorithms are given, i.e., here
sqrt x delivers a list of all integers / naturals / rationals y where y 2 = x.
To this end, the Babylonian method has been adapted by using integerdivisions.
In addition to the precise algorithms, we also provide one approximation algorithm for arbitrary linear ordered fields, where some number y is
computed such that |y 2 − x | < ε.
The major motivation for developing the precise algorithms was given by
CeTA [2], a tool for certifiying termination proofs. Here, non-linear equations
of the form (a1 x1 + . . . an xn )2 = p had to be solved over the integers, where
p is a concrete polynomial. For example, for the equation (ax + by)2 =
2 = 9, and ab = −6,
4x2 − 12xy + 9y 2 one easily figures out √
that a2 = 4, b√
which results in a possible solution a = 4 = 2, b = − 9 = −3.
2
The Babylonian method
The Babylonian method for computing
xi+1 =
n
xi
√
n iteratively computes
+ xi
2
until x2i ≈ n. Note that if x20 ≥ n, then for all i we have both x2i ≥ n and
xi ≥ xi+1 .
3
The Babylonian method using integer division
First, the algorithm is developed for the non-negative integers. Here, the
division operation xy is replaced by x div y = bof-int x / of-int yc. Note
that replacing bof-int x / of-int yc by dof-int x / of-int ye would lead to
non-termination in the following algorithm.
We explicititly develop the algorithm on the integers and not on the
naturals, as the calculations on the integers have been much easier. For
example, y−x+x = y on the integers, which would require the side-condition
y ≥ x for the naturals. These conditions will make the reasoning much more
tedious—as we have experienced in an earlier state of this development where
everything was based on naturals.
For the main algorithm we use the auxiliary guard ¬ (x < 0 ∨ n <
0 ) to ensure that we do not have to worry about negative numbers for the
2
termination argument and during the soundness proof. Moreover, since the
elements x0 , x1 , x2 , . . . are monotone decreasing, we abort as soon as x2i ≤ n.
function sqrt-babylon-int-main :: int ⇒ int ⇒ int option where
sqrt-babylon-int-main x n = (if (x < 0 ∨ n < 0 ) then None else (if x ∗ x ≤ n
then (if x ∗ x = n then Some x else None)
else sqrt-babylon-int-main ((n div x + x ) div 2 ) n))
by pat-completeness auto
For the executable algorithm we omit the guard and use a let-construction
partial-function (tailrec) sqrt-int-main :: int ⇒ int ⇒ int option where
[code]: sqrt-int-main x n = (let x2 = x ∗ x in if x2 ≤ n then (if x2 = n then
Some x else None)
else sqrt-int-main ((n div x + x ) div 2 ) n)
Once we have proven soundness of sqrt-babylon-int-main and equivalence to sqrt-int-main, it is easy to assemble the following algorithm which
computes all square roots for arbitrary integers.
definition sqrt-int :: int ⇒ int list where
sqrt-int x ≡ if x < 0 then [] else case sqrt-int-main x x of Some y ⇒ if y = 0
then [0 ] else [y,−y] | None ⇒ []
For proving soundness, we first need some basic properties of integers.
lemma int-lesseq-square: (z :: int) ≤ z ∗ z
proof (cases z ≥ 0 )
case True
have 1 ∗ z = z 0 ∗ z = 0 by (auto simp: field-simps)
with True show ?thesis by (metis abs-of-nonneg mult-right-mono not-le zabs-less-one-iff )
qed auto
lemma square-int-pos-mono: assumes x : 0 ≤ (x :: int) and y: 0 ≤ y
shows (x ∗ x ≤ y ∗ y) = (x ≤ y)
using assms by (metis mult-mono mult-strict-mono 0 not-less)
lemma square-int-pos-inj : assumes x : 0 ≤ (x :: int) and y: 0 ≤ y
and id : x ∗ x = y ∗ y
shows x = y
proof −
from square-int-pos-mono[OF x y] id have xy: x ≤ y by auto
from square-int-pos-mono[OF y x ] id have yx : y ≤ x by auto
from xy yx show ?thesis by auto
qed
lemma mod-div-equality-int: (n :: int) div x ∗ x = n − n mod x
using mod-div-equality[of n x ] by arith
lemma iteration-mono-eq: assumes xn: x ∗ x = (n :: int)
shows (n div x + x ) div 2 = x unfolding xn[symmetric]
by simp
3
The following property is the essential property for proving termination
of sqrt-babylon-int-main.
lemma iteration-mono-less: assumes x : x ≥ 0
and n: n ≥ 0
and xn: x ∗ x > (n :: int)
shows (n div x + x ) div 2 < x
proof −
let ?sx = (n div x + x ) div 2
from xn have xn-le: x ∗ x ≥ n by auto
from xn x n have x0 : x > 0 by (cases x = 0 , auto)
have n div x ∗ x ≤ n unfolding mod-div-equality-int
using transfer-nat-int-function-closures x n
by simp
also have . . . ≤ x ∗ x using xn by auto
finally have le: n div x ≤ x using x x0 by auto
have ?sx ≤ ((x ∗ x ) div x + x ) div 2
by (rule zdiv-mono1 , insert le, auto)
also have . . . ≤ x by fastforce
finally have le: ?sx ≤ x .
{
assume ?sx = x
hence n div x ≥ x by auto
from mult-right-mono[OF this x ]
have ge: n div x ∗ x ≥ x ∗ x .
from mod-div-equality[of n x ] have n div x ∗ x = n − n mod x by arith
from ge[unfolded this]
have le: x ∗ x ≤ n − n mod x .
have ge: n mod x ≥ 0 using n x transfer-nat-int-function-closures by auto
from le ge
have n ≥ x ∗ x by auto
with xn have False by auto
}
with le show ?thesis by fastforce
qed
lemma iteration-mono-lesseq: assumes x : x ≥ 0 and n: n ≥ 0 and xn: x ∗ x ≥
(n :: int)
shows (n div x + x ) div 2 ≤ x
proof (cases x ∗ x = n)
case True
from iteration-mono-eq[OF this] show ?thesis by simp
next
case False
with assms have x ∗ x > n by auto
from iteration-mono-less[OF x n this]
show ?thesis by simp
qed
termination
4
proof −
let ?mm = λ x n :: int. nat x
let ?m1 = λ (x ,n). ?mm x n
let ?m = measures [?m1 ]
show ?thesis
proof (relation ?m, simp)
fix x n :: int
assume ¬ x ∗ x ≤ n
hence x : x ∗ x > n by auto
assume ¬ (x < 0 ∨ n < 0 )
hence x-n: x ≥ 0 n ≥ 0 by auto
from x x-n have x0 : x > 0 by (cases x = 0 , auto)
from iteration-mono-less[OF x-n x ] x0
show (((n div x + x ) div 2 ,n),(x ,n)) ∈ ?m by auto
qed
qed
We next prove that sqrt-int-main is a correct implementation of sqrt-babylon-int-main.
We additionally prove that the result is always positive, which poses no overhead and allows to share the inductive proof.
lemma sqrt-int-main-babylon-pos: x ≥ 0 =⇒ n ≥ 0 =⇒ sqrt-int-main x n =
sqrt-babylon-int-main x n ∧ (sqrt-int-main x n = Some y −→ y ≥ 0 )
proof (induct x n rule: sqrt-babylon-int-main.induct)
case (1 x n)
hence id : (x < 0 ∨ n < 0 ) = False by auto
note d = sqrt-int-main.simps[of x n] sqrt-babylon-int-main.simps[of x n] id
if-False Let-def
show ?case
proof (cases x ∗ x ≤ n)
case True
thus ?thesis unfolding d using 1 (2 ) by auto
next
case False
hence id : (x ∗ x ≤ n) = False by simp
from 1 (3 ) 1 (2 ) have not: ¬ (x < 0 ∨ n < 0 ) by auto
show ?thesis unfolding d id
by (rule 1 (1 )[OF not False - 1 (3 )], insert transfer-nat-int-function-closures
1 (2 ) 1 (3 ), auto)
qed
qed
lemma sqrt-int-main: x ≥ 0 =⇒ n ≥ 0 =⇒ sqrt-int-main x n = sqrt-babylon-int-main
xn
using sqrt-int-main-babylon-pos by blast
lemma sqrt-int-main-pos: x ≥ 0 =⇒ n ≥ 0 =⇒ sqrt-int-main x n = Some y =⇒
y ≥0
using sqrt-int-main-babylon-pos by blast
Soundness of sqrt-babylon-int-main is trivial.
5
lemma sqrt-babylon-int-main-sound : sqrt-babylon-int-main x n = Some y =⇒ y ∗
y =n
proof (induct x n rule: sqrt-babylon-int-main.induct)
case (1 x n)
from 1 (2 ) have not: ¬ (x < 0 ∨ n < 0 ) by (cases x < 0 ∨ n < 0 , auto)
show ?case
proof (cases x ∗ x ≤ n)
case False
with 1 (1 )[OF not this] not 1 (2 ) show ?thesis by simp
next
case True
thus ?thesis using 1 (2 ) not by (cases x ∗ x = n, auto)
qed
qed
For completeness, we first need the following crucial inquality, which
would be trivial if one would use standard division instead of integer division.
(Proving this equation has required a significant amount of time during the
development, as we first made several proof attemps which have been dead
ends.)
lemma square-inequality: 2 ∗ x ∗ y ≤ x ∗ x div y ∗ y + y ∗ (y :: int)
proof −
have pos: 0 ≤ (y − x ) ∗ (y − x ) div y ∗ y
by (metis transfer-nat-int-function-closures div-0 eq-iff linear neg-imp-zdiv-nonneg-iff
not-le split-mult-pos-le)
have x ∗ x div y ∗ y + y ∗ y =
((y − x ) ∗ (y − x ) + y∗(2 ∗x − y)) div y ∗ y + y ∗ y by (simp add : field-simps)
also have . . . = 2 ∗y∗x − y∗y + (y−x )∗(y−x ) div y ∗ y + y ∗ y
proof (cases y = 0 )
case False
show ?thesis unfolding div-mult-self2 [OF False] by (auto simp: field-simps)
qed simp
also have . . . = 2 ∗x ∗y + (y−x )∗(y−x ) div y ∗ y by (auto simp: field-simps)
also have . . . ≥ 2 ∗x ∗y using pos by auto
finally show ?thesis .
qed
lemma sqrt-babylon-int-main-complete: assumes x0 : x ≥ 0
shows x ∗ x = n =⇒ y ∗ y ≥ n =⇒ y ≥ 0 =⇒ n ≥ 0 =⇒ sqrt-babylon-int-main
y n = Some x
proof (induct y n rule: sqrt-babylon-int-main.induct)
case (1 y n)
have nx : n = x ∗ x and xn: x ∗ x = n using 1 (2 ) by auto
from 1 (4 ) have y0 : y ≥ 0 .
from 1 (5 ) have n0 : n ≥ 0 .
let ?sqrt = sqrt-babylon-int-main
from y0 n0 have not: ¬ (y < 0 ∨ n < 0 ) by auto
show ?case
proof (cases y ∗ y ≤ n)
6
case False
let ?sy = (n div y + y) div 2
let ?sx = (n div x + x ) div 2
from False y0 n0 have id : ?sqrt y n = ?sqrt ?sy n by auto
from False xn have x ∗ x ≤ y ∗ y by auto
from this[unfolded square-int-pos-mono[OF x0 y0 ]] have xy: x ≤ y by auto
have sx0 : ?sx ≥ 0 using transfer-nat-int-function-closures n0 x0 by auto
have sy0 : ?sy ≥ 0 using transfer-nat-int-function-closures n0 y0 by auto
show ?thesis unfolding id
proof (rule 1 (1 )[OF not False xn - sy0 n0 ])
have n = ?sx ∗ ?sx unfolding nx by simp
also have . . . ≤ ?sy ∗ ?sy unfolding square-int-pos-mono[OF sx0 sy0 ]
proof −
have ?sx = 2 ∗ x div 2 unfolding nx by simp
also have . . . ≤ ?sy
proof (rule zdiv-mono1 )
show 2 ∗ x ≤ n div y + y
proof (cases y = 0 )
case True
thus ?thesis using xy by auto
next
case False
with y0 have y: y > 0 by auto
show ?thesis
proof (rule mult-right-le-imp-le[OF - y], unfold nx )
show 2 ∗ x ∗ y ≤ (x ∗ x div y + y) ∗ y using square-inequality by
(auto simp: field-simps)
qed
qed
qed simp
finally show ?sx ≤ ?sy .
qed
finally
show n ≤ ?sy ∗ ?sy .
qed
next
case True
with 1 (3 ) have yn: y ∗ y = n by auto
with nx have x ∗ x = y ∗ y by auto
from square-int-pos-inj [OF x0 y0 this] yn not show ?thesis by auto
qed
qed
Having proven soundness and completeness of sqrt-babylon-int-main, it
is easy to prove soundness of sqrt-int.
lemma sqrt-int[simp]: set (sqrt-int x ) = {y. y ∗ y = x }
proof −
note d = sqrt-int-def
show ?thesis
7
proof (cases x < 0 ∨ sqrt-int-main x x = None)
case True
hence left: set (sqrt-int x ) = {} unfolding d by (cases x < 0 , auto)
{
fix y
assume x : y ∗ y = x
have ∃ y. y ∗ y = x ∧ y ≥ 0
proof (cases y ≥ 0 )
case True
with x show ?thesis by auto
next
case False
show ?thesis
by (rule exI [of - −y], insert False x , auto)
qed
then obtain y where x : y ∗ y = x and y: y ≥ 0 by auto
from y x have x0 : x ≥ 0 by auto
from sqrt-babylon-int-main-complete[OF y x int-lesseq-square x0 x0 ] True
have False unfolding sqrt-int-main[OF x0 x0 ] by auto
}
with left show ?thesis by auto
next
case False
then obtain y where Some: sqrt-int-main x x = Some y and x : x ≥ 0 by
auto
hence left: set (sqrt-int x ) = {y,−y} unfolding d by (cases y = 0 , auto)
have ∃ z . z ≥ 0 ∧ {y,−y} = {z ,−z }
proof (cases y ≥ 0 )
case False
show ?thesis by (rule exI [of - −y], insert False, auto)
qed auto
then obtain z where z : z ≥ 0 and yz : {y,−y} = {z ,−z } by auto
from sqrt-babylon-int-main-sound [OF Some[unfolded sqrt-int-main[OF x x ]]]
have y: y ∗ y = x −y ∗ −y = x by auto
with yz have zz : z ∗ z = x −z ∗ −z = x by auto
from y have set (sqrt-int x ) ⊆ {y. y ∗ y = x } unfolding left by auto
moreover
{
fix u
assume u: u ∗ u = x
have u ∈ {z ,−z }
proof (cases u ≥ 0 )
case True
from u zz have u ∗ u = z ∗ z by auto
from square-int-pos-inj [OF True z this]
show ?thesis by auto
next
case False
hence up: −u ≥ 0 by auto
8
from u zz have −u ∗ −u = z ∗ z by auto
from square-int-pos-inj [OF up z this] show ?thesis by auto
qed
}
hence {y. y ∗ y = x } ⊆ set (sqrt-int x ) unfolding left yz by auto
finally show ?thesis by auto
qed
qed
4
Square roots for the naturals
The idea is to use the algorithm for the integers and then convert between
naturals and integers.
In the following lemma, we first observe that the first result of sqrt-int
is always non-negative.
lemma sqrt-int-empty-0-pos-neg: ∃ y. y > 0 ∧ (sqrt-int x = [] ∨ sqrt-int x = [0 ]
∨ sqrt-int x = [y,−y])
proof (cases sqrt-int x = [])
case True
show ?thesis
by (rule exI [of - 1 ], insert True, auto)
next
case False
note d = sqrt-int-def
from False have x : x ≥ 0 unfolding d by (cases x ≥ 0 , auto)
from False x obtain y where main: sqrt-int-main x x = Some y unfolding d
by (cases sqrt-int-main x x , auto)
from sqrt-int-main-pos[OF x x main] have y: y ≥ 0 by auto
show ?thesis
proof (cases y = 0 )
case True
with main x have res: sqrt-int x = [0 ] unfolding d by auto
show ?thesis
by (rule exI [of - 1 ], insert True res, auto)
next
case False
with y have y: y > 0 by auto
with main x have res: sqrt-int x = [y,−y] unfolding d by auto
thus ?thesis using y by auto
qed
qed
With this knowledge, it is easy to define a square-root function on the
naturals.
definition sqrt-nat :: nat ⇒ nat list
where sqrt-nat x ≡ case (sqrt-int (int x )) of [] ⇒ [] | x # xs ⇒ [nat x ]
The soundness of sqrt-nat is straight-forward.
9
lemma sqrt-nat[simp]: set (sqrt-nat x ) = { y. y ∗ y = x } (is ?l = ?r )
proof −
note d = sqrt-nat-def
let ?sx = sqrt-int (int x )
from sqrt-int-empty-0-pos-neg[of int x ] obtain y where y: y > 0
and choice:V?sx = [] ∨ (?sx = [0 ] ∨ ?sx = [y,−y]) by auto
have arith:
n. of-nat (n ∗ n) = of-nat n ∗ of-nat n using int-arith-rules by
auto
{
fix z
assume z ∈ ?l
with choice have hd : hd ?sx ∈ set (?sx ) ∧ z = nat (hd ?sx ) ∧ ?sx 6= []
unfolding d
by (cases ?sx , auto)
hence z : z = nat (hd ?sx ) by auto
from hd have hd ?sx ∗ hd ?sx = int x hd ?sx ≥ 0 using y choice by auto
hence z ∗ z = x unfolding z by (metis arith int-nat-eq nat-int)
hence z ∈ ?r by auto
}
moreover
{
fix z
assume z ∈ ?r
hence x : x = z ∗ z by auto
from arg-cong[OF this, of int]
have x : int x = int z ∗ int z unfolding arith .
hence mem: int z ∈ set ?sx by auto
from choice mem have ?sx = [0 ] ∨ ?sx = [y,−y] by (cases ?sx , auto)
hence z ∈ ?l
proof
assume ?sx = [0 ]
with mem have z : z = 0 by auto
with x have x : x = 0 by auto
thus ?thesis unfolding d z x by eval
next
assume sx : ?sx = [y,−y]
with mem y have y: y = int z by auto
show ?thesis unfolding d sx list.simps y by auto
qed
}
ultimately show ?thesis by blast
qed
5
Square roots for the rationals
For the rationals, the idea again, to apply the algorithm for the integers and
then convert between integers and rationals. Here, the essential idea is to
compute the square-roots of the numerator and denominator separately.
10
definition sqrt-rat :: rat ⇒ rat list
where sqrt-rat x ≡ case quotient-of x of (z ,n) ⇒ (case sqrt-int n of
[] ⇒ []
| sn # xs ⇒ map (λ sz . of-int sz / of-int sn) (sqrt-int z ))
Whereas soundness of sqrt-rat is simple, it is a bit more tedious to show
that all roots are computed, which uses facts on coprime.
lemma sqrt-rat[simp]: set (sqrt-rat x ) = { y. y ∗ y = x } (is ?l = ?r )
proof −
note d = sqrt-rat-def
obtain z n where q: quotient-of x = (z ,n) by force
let ?sz = sqrt-int z
let ?sn = sqrt-int n
from q have n: n > 0 by (rule quotient-of-denom-pos)
from q have x : x = of-int z / of-int n by (rule quotient-of-div )
{
fix y
assume y ∈ ?l
note y = this[unfolded d q split]
then obtain sn snn where sn: ?sn = sn # snn by (cases ?sn, auto)
from y[unfolded sn] obtain sz where sz : sz ∈ set ?sz and y: y = of-int sz /
of-int sn by auto
from sn have sn: sn ∈ set ?sn by auto
from sn have sn: n = sn ∗ sn by auto
from sz have sz : z = sz ∗ sz by auto
have y ∈ ?r unfolding x sn sz y by auto
}
moreover
{
fix y
assume y ∈ ?r
obtain yz yn where yq: quotient-of y = (yz ,yn) by force
let ?syz = (of-int yz ) :: rat
let ?syn = (of-int yn) :: rat
from yq have yn: yn > 0 by (rule quotient-of-denom-pos)
from yq have y: y = ?syz / ?syn by (rule quotient-of-div )
from hy ∈ ?r i x
have of-int z / of-int n = y ∗ y by simp
also have . . . = ?syz ∗ ?syz / (?syn ∗ ?syn) unfolding y by simp
finally haveVid : of-int z / of-int n = ?syz ∗ ?syz / (?syn ∗ ?syn) .
have arith:
x y. of-int x ∗ of-int y = of-int (x ∗ y) by simp
have sn: n = yn ∗ yn
proof −
from quotient-of-coprime[OF q] have cop: coprime z n .
from quotient-of-coprime[OF yq] have ycop: coprime yz yn .
hence ycop: coprime (yz ∗ yz ) (yn ∗ yn) by (metis coprime-mul-eq-int
gcd-mult-cancel-int)
from arg-cong[OF id , of λ x . x ∗ of-int n ∗ ?syn ∗ ?syn]
have of-int z ∗ (?syn ∗ ?syn) = (?syz ∗ ?syz ) ∗ of-int n using n yn by (simp
11
add : field-simps)
from this[unfolded arith] have id : z ∗ (yn ∗ yn) = (yz ∗ yz ) ∗ n by linarith
hence abs z ∗ abs (yn ∗ yn) = abs (yz ∗ yz ) ∗ abs n unfolding abs-mult[symmetric]
by simp
with coprime-crossproduct-int[OF cop ycop] have abs n = abs (yn ∗ yn) by
auto
with yn n show n = yn ∗ yn by auto
qed
with id n have sz : of-int z = ?syz ∗ ?syz by auto
from sz have sz : z = yz ∗ yz unfolding arith by linarith
from sqrt-int-empty-0-pos-neg[of n] obtain sn where
s-n: ?sn = [] ∨ ?sn = [0 ] ∨ ?sn = [sn,−sn] and pos: sn > 0 by auto
from sn have y-n: yn ∈ set ?sn by simp
with s-n yn have s-n: ?sn = [sn,−sn] by (cases ?sn, auto)
from y-n pos s-n yn have sn: sn = yn by auto
with s-n have sn: ?sn = [yn,−yn] by auto
from sz have yz : yz ∈ set ?sz by auto
have y ∈ ?l unfolding d q using sn y yz by auto
}
ultimately show ?thesis by auto
qed
6
Approximating square roots
The difference to the previous algorithms is that now we abort, once the
distance is below . Moreover, here we use standard division and not integer
division.
We first provide the executable version without guard (0 :: 0a) < x as partial function, and afterwards prove termination and soundness for a similar
algorithm that is defined within the upcoming locale.
partial-function (tailrec) sqrt-approx-main-impl :: 0a :: linordered-field ⇒ 0a ⇒
0
a ⇒ 0a where
[code]: sqrt-approx-main-impl ε n x = (if x ∗ x − n < ε then x else sqrt-approx-main-impl
εn
((n / x + x ) / 2 ))
We setup a locale where we ensure that we have standard assumptions:
positive and positive n. We require sort floor-ceiling, since bx c is used for
the termination argument.
locale sqrt-approximation =
fixes ε :: 0a :: {linordered-field ,floor-ceiling}
and n :: 0a
assumes ε : ε > 0
and n: n > 0
begin
function sqrt-approx-main :: 0a ⇒ 0a where
12
sqrt-approx-main x = (if x > 0 then (if x ∗ x − n < ε then x else sqrt-approx-main
((n / x + x ) / 2 )) else 0 )
by pat-completeness auto
Termination essentially is a proof of convergence. Here, one complication
is the fact that the limit is not always defined. E.g., if 0a is rat then there
is no square root of 2. Therefore, the error-rate √xn − 1 is not expressible.
Instead we use the expression
any square-root operation.
x2
n
− 1 as error-rate which does not require
termination
proof −
def er ≡ λ x . (x ∗ x / n − 1 )
def c ≡ 2 ∗ n / ε
def m ≡ λ x . nat b c ∗ er x c
have c: c > 0 unfolding c-def using n ε
by (auto intro: divide-pos-pos)
show ?thesis
proof
show wf (measures [m]) by simp
next
fix x
assume x : 0 < x and xe: ¬ x ∗ x − n < ε
def y ≡ (n / x + x ) / 2
show ((n / x + x ) / 2 ,x ) ∈ measures [m] unfolding y-def [symmetric]
proof (rule measures-less)
from n have inv-n: 1 / n > 0 by (auto intro: divide-pos-pos)
from xe have x ∗ x − n ≥ ε by simp
from this[unfolded mult-le-cancel-left-pos[OF inv-n, of ε, symmetric]]
have erxen: er x ≥ ε / n unfolding er-def using n by (simp add : field-simps)
have en: ε / n > 0 and ne: n / ε > 0 using ε n by (auto intro: divide-pos-pos)
from en erxen have erx : er x > 0 by linarith
have pos: er x ∗ 4 + er x ∗ (er x ∗ 4 ) > 0 using erx
by (auto intro: add-pos-nonneg)
have er y = 1 / 4 ∗ (n / (x ∗ x ) − 2 + x ∗ x / n) unfolding er-def y-def
using x n
by (simp add : field-simps)
also have . . . = 1 / 4 ∗ er x ∗ er x / (1 + er x ) unfolding er-def using x
n
by (simp add : field-simps)
finally have er y = 1 / 4 ∗ er x ∗ er x / (1 + er x ) .
also have . . . < 1 / 4 ∗ (1 + er x ) ∗ er x / (1 + er x ) using erx erx pos
by (auto intro: mult-pos-pos simp: field-simps)
also have . . . = er x / 4 using erx by (simp add : field-simps)
finally have er-y-x : er y ≤ er x / 4 by linarith
from erxen have c ∗ er x ≥ 2 unfolding c-def mult-le-cancel-left-pos[OF
ne, of - er x , symmetric]
using n ε by (auto simp: field-simps)
hence pos: bc ∗ er x c > 0 bc ∗ er x c ≥ 2 by auto
13
show m y < m x unfolding m-def nat-mono-iff [OF pos(1 )]
proof −
have bc ∗ er yc ≤ bc ∗ (er x / 4 )c
by (rule floor-mono, unfold mult-le-cancel-left-pos[OF c], rule er-y-x )
also have . . . < bc ∗ er x / 4 + 1 c by auto
also have . . . ≤ bc ∗ er x c
by (rule floor-mono, insert pos(2 ), simp add : field-simps)
finally show bc ∗ er yc < bc ∗ er x c .
qed
qed
qed
qed
Once termination is proven, it is easy to show equivalence of sqrt-approx-main-impl
and sqrt-approx-main.
lemma sqrt-approx-main-impl : x > 0 =⇒ sqrt-approx-main-impl ε n x = sqrt-approx-main
x
proof (induct x rule: sqrt-approx-main.induct)
case (1 x )
hence x : x > 0 by auto
hence nx : 0 < (n / x + x ) / 2 using n by (auto intro: divide-pos-pos pos-add-strict)
note simps = sqrt-approx-main-impl .simps[of - - x ] sqrt-approx-main.simps[of
x]
show ?case
proof (cases x ∗ x − n < ε)
case True
thus ?thesis unfolding simps using x by auto
next
case False
show ?thesis using 1 (1 )[OF x False nx ] unfolding simps using x False by
auto
qed
qed
Also soundness is not complicated.
lemma sqrt-approx-main-sound : assumes x : x > 0 and xx : x ∗ x > n
shows sqrt-approx-main x ∗ sqrt-approx-main x > n ∧ sqrt-approx-main x ∗
sqrt-approx-main x − n < ε
using assms
proof (induct x rule: sqrt-approx-main.induct)
case (1 x )
from 1 have x : x > 0 (x > 0 ) = True by auto
note simp = sqrt-approx-main.simps[of x , unfolded x if-True]
show ?case
proof (cases x ∗ x − n < ε)
case True
with 1 show ?thesis unfolding simp by simp
next
case False
14
let ?y = (n / x + x ) / 2
from False simp have simp: sqrt-approx-main x = sqrt-approx-main ?y by
simp
from n x have y: ?y > 0 by (auto intro: divide-pos-pos pos-add-strict)
note IH = 1 (1 )[OF x (1 ) False y]
from x have x4 : 4 ∗ x ∗ x > 0 by (auto intro: mult-sign-intros)
show ?thesis unfolding simp
proof (rule IH )
show n < ?y ∗ ?y
unfolding mult-less-cancel-left-pos[OF x4 , of n, symmetric]
proof −
have id : 4 ∗ x ∗ x ∗ (?y ∗ ?y) = 4 ∗ x ∗ x ∗ n + (n − x ∗ x ) ∗ (n − x ∗
x ) using x (1 )
by (simp add : field-simps)
from 1 (3 ) have x ∗ x − n > 0 by auto
from mult-pos-pos[OF this this]
show 4 ∗ x ∗ x ∗ n < 4 ∗ x ∗ x ∗ (?y ∗ ?y) unfolding id
by (simp add : field-simps)
qed
qed
qed
qed
end
It remains to assemble everything into one algorithm.
definition sqrt-approx :: 0a :: {linordered-field ,floor-ceiling} ⇒ 0a ⇒ 0a where
sqrt-approx ε x ≡ if ε > 0 then (if x = 0 then 0 else let xpos = abs x in
sqrt-approx-main-impl ε xpos (xpos + 1 )) else 0
lemma sqrt-approx : assumes ε: ε > 0
shows |sqrt-approx ε x ∗ sqrt-approx ε x − |x || < ε
proof (cases x = 0 )
case True
with ε show ?thesis unfolding sqrt-approx-def by auto
next
case False
let ?x = |x |
let ?sqrti = sqrt-approx-main-impl ε ?x (?x + 1 )
let ?sqrt = sqrt-approximation.sqrt-approx-main ε ?x (?x + 1 )
def sqrt ≡ ?sqrt
from False have x : ?x > 0 ?x + 1 > 0 by auto
interpret sqrt-approximation ε ?x
by (unfold-locales, insert x ε, auto)
from False ε have sqrt-approx ε x = ?sqrti unfolding sqrt-approx-def by (simp
add : Let-def )
also have ?sqrti = ?sqrt
by (rule sqrt-approx-main-impl , auto)
15
finally have id : sqrt-approx ε x = sqrt unfolding sqrt-def .
have sqrt: sqrt ∗ sqrt > ?x ∧ sqrt ∗ sqrt − ?x < ε unfolding sqrt-def
by (rule sqrt-approx-main-sound [OF x (2 )], insert x mult-pos-pos[OF x (1 ) x (1 )],
auto simp: field-simps)
show ?thesis unfolding id using sqrt by auto
qed
7
Some tests
Testing executabity and show that sqrt 2 is irrational
lemma ¬ (∃ i :: rat. i ∗ i = 2 )
proof −
have set (sqrt-rat 2 ) = {} by eval
thus ?thesis by simp
qed
Testing speed
lemma ¬ (∃ i :: int. i ∗ i = 1234567890123456789012345678901234567890 )
proof −
have set (sqrt-int 1234567890123456789012345678901234567890 ) = {} by eval
thus ?thesis by simp
qed
The following test
value let ε = 1 / 100000000 :: rat; s = sqrt-approx ε 2 in (s, s ∗ s − 2 , |s ∗ s
− 2 | < ε)
results in (1.4142135623731116, 4.738200762148612e-14, True).
end
References
[1] T. Heath. A History of Greek Mathematics, volume 2, pages 323–326.
Clarendon Press, 1921.
[2] R. Thiemann and C. Sternagel. Certification of termination proofs using
CeTA. In Proc. TPHOLs’09, volume 5674 of LNCS, pages 452–468, 2009.
16