79
ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties of Pi ...................... 2 1.2 Composition With a Restricted Domain: compose ...... 2 1.3 Bounded Abstraction: restrict .................. 3 1.4 Extensionality ........................... 3 2 Groups 5 3 From Magmas to Groups 5 3.1 Definitions ............................. 5 3.2 Cancellation Laws and Basic Properties ............ 11 3.3 Substructures ........................... 12 3.4 Direct Products .......................... 14 3.5 Homomorphisms ......................... 16 3.6 Commutative Structures ..................... 18 3.7 Definition ............................. 18 4 Product Operator 19 4.1 Left-commutative operations .................. 20 4.2 Commutative monoids ...................... 24 4.3 Products over Finite Sets .................... 25 5 Extracting code from tuples definition and operations 29 5.1 Definition of complections .................... 29 5.2 Definition of tuples ........................ 30 5.3 Lemmas about composition of tuples .............. 31 5.4 Code extraction .......................... 32 6 The Greatest Common Divisor and Euclid’s algorithm Christophe Tabacznyj and Lawrence C Paulson 32 1

ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

  • Upload
    others

  • View
    0

  • Download
    0

Embed Size (px)

Citation preview

Page 1: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

ProgramExtract

Jesus Maria Aransay

October 12, 2005

Contents

1 Pi and Function Sets Florian Kammueller and Lawrence CPaulson 11.1 Basic Properties of Pi . . . . . . . . . . . . . . . . . . . . . . 21.2 Composition With a Restricted Domain: compose . . . . . . 21.3 Bounded Abstraction: restrict . . . . . . . . . . . . . . . . . . 31.4 Extensionality . . . . . . . . . . . . . . . . . . . . . . . . . . . 3

2 Groups 5

3 From Magmas to Groups 53.1 Definitions . . . . . . . . . . . . . . . . . . . . . . . . . . . . . 53.2 Cancellation Laws and Basic Properties . . . . . . . . . . . . 113.3 Substructures . . . . . . . . . . . . . . . . . . . . . . . . . . . 123.4 Direct Products . . . . . . . . . . . . . . . . . . . . . . . . . . 143.5 Homomorphisms . . . . . . . . . . . . . . . . . . . . . . . . . 163.6 Commutative Structures . . . . . . . . . . . . . . . . . . . . . 183.7 Definition . . . . . . . . . . . . . . . . . . . . . . . . . . . . . 18

4 Product Operator 194.1 Left-commutative operations . . . . . . . . . . . . . . . . . . 204.2 Commutative monoids . . . . . . . . . . . . . . . . . . . . . . 244.3 Products over Finite Sets . . . . . . . . . . . . . . . . . . . . 25

5 Extracting code from tuples definition and operations 295.1 Definition of complections . . . . . . . . . . . . . . . . . . . . 295.2 Definition of tuples . . . . . . . . . . . . . . . . . . . . . . . . 305.3 Lemmas about composition of tuples . . . . . . . . . . . . . . 315.4 Code extraction . . . . . . . . . . . . . . . . . . . . . . . . . . 32

6 The Greatest Common Divisor and Euclid’s algorithm ChristopheTabacznyj and Lawrence C Paulson 32

1

Page 2: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

7 Permutations Lawrence C Paulson and Thomas M Rasmussen 377.1 Some examples of rule induction on permutations . . . . . . . 377.2 Ways of making new permutations . . . . . . . . . . . . . . . 387.3 Further results . . . . . . . . . . . . . . . . . . . . . . . . . . 397.4 Removing elements . . . . . . . . . . . . . . . . . . . . . . . . 39

8 Fundamental Theorem of Arithmetic (unique factorizationinto primes) 408.1 Definitions . . . . . . . . . . . . . . . . . . . . . . . . . . . . . 408.2 Arithmetic . . . . . . . . . . . . . . . . . . . . . . . . . . . . . 418.3 Prime list and product . . . . . . . . . . . . . . . . . . . . . . 418.4 Sorting . . . . . . . . . . . . . . . . . . . . . . . . . . . . . . . 438.5 Permutation . . . . . . . . . . . . . . . . . . . . . . . . . . . . 448.6 Existence . . . . . . . . . . . . . . . . . . . . . . . . . . . . . 448.7 Uniqueness . . . . . . . . . . . . . . . . . . . . . . . . . . . . 45

9 The set of prime numbers is infinite 479.1 Definition of factorial and basic facts . . . . . . . . . . . . . . 479.2 New definition of division for naturals . . . . . . . . . . . . . 489.3 The infinity of the set of primes using Isabelle definitions . . 509.4 Euclid’s lemma and the set of primes is infinite . . . . . . . . 539.5 Recursive definition of division . . . . . . . . . . . . . . . . . 539.6 Equivalence between the Isabelle defined dvd and the recur-

sive definition . . . . . . . . . . . . . . . . . . . . . . . . . . . 549.7 Definition of prime number recursively . . . . . . . . . . . . . 579.8 Equivalence between the Isabelle prime definition and our

recursive one . . . . . . . . . . . . . . . . . . . . . . . . . . . 589.9 Function yielding a prime number smaller than the argument

and dividing it . . . . . . . . . . . . . . . . . . . . . . . . . . 679.10 Properties of someprimedivisor . . . . . . . . . . . . . . . . . 709.11 Euclid’s lemma in a constructive way . . . . . . . . . . . . . . 769.12 Code extraction . . . . . . . . . . . . . . . . . . . . . . . . . . 76

10 Definition of monoid of integers and the rev of a monoid 7710.1 Code extraction . . . . . . . . . . . . . . . . . . . . . . . . . . 77

1 Pi and Function Sets Florian Kammueller andLawrence C Paulson

theory FuncSet = Main:

constdefsPi :: [ ′a set , ′a => ′b set ] => ( ′a => ′b) set

Pi A B == {f . ∀ x . x ∈ A −−> f (x ) ∈ B(x )}

2

Page 3: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

extensional :: ′a set => ( ′a => ′b) setextensional A == {f . ∀ x . x∼:A −−> f (x ) = arbitrary}

restrict :: [ ′a => ′b, ′a set ] => ( ′a => ′b)restrict f A == (%x . if x ∈ A then f x else arbitrary)

syntax@Pi :: [pttrn, ′a set , ′b set ] => ( ′a => ′b) set ((3PI -:-./ -) 10 )funcset :: [ ′a set , ′b set ] => ( ′a => ′b) set (infixr −> 60 )@lam :: [pttrn, ′a set , ′a => ′b] => ( ′a=> ′b) ((3%-:-./ -) [0 ,0 ,3 ] 3 )

syntax (xsymbols)@Pi :: [pttrn, ′a set , ′b set ] => ( ′a => ′b) set ((3Π -∈-./ -) 10 )funcset :: [ ′a set , ′b set ] => ( ′a => ′b) set (infixr → 60 )@lam :: [pttrn, ′a set , ′a => ′b] => ( ′a=> ′b) ((3λ-∈-./ -) [0 ,0 ,3 ] 3 )

syntax (HTML output)@Pi :: [pttrn, ′a set , ′b set ] => ( ′a => ′b) set ((3Π -∈-./ -) 10 )@lam :: [pttrn, ′a set , ′a => ′b] => ( ′a=> ′b) ((3λ-∈-./ -) [0 ,0 ,3 ] 3 )

translationsPI x :A. B => Pi A (%x . B)A −> B => Pi A (-K B)%x :A. f == restrict (%x . f ) A

constdefscompose :: [ ′a set , ′b => ′c, ′a => ′b] => ( ′a => ′c)compose A g f == λx∈A. g (f x )

print-translation {∗ [(Pi , dependent-tr ′ (@Pi , funcset))] ∗}

1.1 Basic Properties of Pi

lemma Pi-I : (!!x . x ∈ A ==> f x ∈ B x ) ==> f ∈ Pi A Bby (simp add : Pi-def )

lemma funcsetI : (!!x . x ∈ A ==> f x ∈ B) ==> f ∈ A −> Bby (simp add : Pi-def )

lemma Pi-mem: [|f : Pi A B ; x ∈ A|] ==> f x ∈ B xby (simp add : Pi-def )

lemma funcset-mem: [|f ∈ A −> B ; x ∈ A|] ==> f x ∈ Bby (simp add : Pi-def )

lemma Pi-eq-empty : ((PI x : A. B x ) = {}) = (∃ x∈A. B(x ) = {})apply (simp add : Pi-def , auto)

Converse direction requires Axiom of Choice to exhibit a function picking an ele-

3

Page 4: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

ment from each non-empty B x

apply (drule-tac x = %u. SOME y . y ∈ B u in spec, auto)apply (cut-tac P= %y . y ∈ B x in some-eq-ex , auto)done

lemma Pi-empty [simp]: Pi {} B = UNIVby (simp add : Pi-def )

lemma Pi-UNIV [simp]: A −> UNIV = UNIVby (simp add : Pi-def )

Covariance of Pi-sets in their second argument

lemma Pi-mono: (!!x . x ∈ A ==> B x <= C x ) ==> Pi A B <= Pi A Cby (simp add : Pi-def , blast)

Contravariance of Pi-sets in their first argument

lemma Pi-anti-mono: A ′ <= A ==> Pi A B <= Pi A ′ Bby (simp add : Pi-def , blast)

1.2 Composition With a Restricted Domain: compose

lemma funcset-compose:[| f ∈ A −> B ; g ∈ B −> C |]==> compose A g f ∈ A −> C

by (simp add : Pi-def compose-def restrict-def )

lemma compose-assoc:[| f ∈ A −> B ; g ∈ B −> C ; h ∈ C −> D |]==> compose A h (compose A g f ) = compose A (compose B h g) f

by (simp add : expand-fun-eq Pi-def compose-def restrict-def )

lemma compose-eq : x ∈ A ==> compose A g f x = g(f (x ))by (simp add : compose-def restrict-def )

lemma surj-compose: [| f ‘ A = B ; g ‘ B = C |] ==> compose A g f ‘ A = Cby (auto simp add : image-def compose-eq)

lemma inj-on-compose:[| f ‘ A = B ; inj-on f A; inj-on g B |] ==> inj-on (compose A g f ) A

by (auto simp add : inj-on-def compose-eq)

1.3 Bounded Abstraction: restrict

lemma restrict-in-funcset : (!!x . x ∈ A ==> f x ∈ B) ==> (λx∈A. f x ) ∈ A −>Bby (simp add : Pi-def restrict-def )

lemma restrictI : (!!x . x ∈ A ==> f x ∈ B x ) ==> (λx∈A. f x ) ∈ Pi A Bby (simp add : Pi-def restrict-def )

4

Page 5: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma restrict-apply [simp]:(λy∈A. f y) x = (if x ∈ A then f x else arbitrary)

by (simp add : restrict-def )

lemma restrict-ext :(!!x . x ∈ A ==> f x = g x ) ==> (λx∈A. f x ) = (λx∈A. g x )

by (simp add : expand-fun-eq Pi-def Pi-def restrict-def )

lemma inj-on-restrict-eq : inj-on (restrict f A) A = inj-on f Aby (simp add : inj-on-def restrict-def )

lemma Id-compose:[|f ∈ A −> B ; f ∈ extensional A|] ==> compose A (λy∈B . y) f = f

by (auto simp add : expand-fun-eq compose-def extensional-def Pi-def )

lemma compose-Id :[|g ∈ A −> B ; g ∈ extensional A|] ==> compose A g (λx∈A. x ) = g

by (auto simp add : expand-fun-eq compose-def extensional-def Pi-def )

1.4 Extensionality

lemma extensional-arb: [|f ∈ extensional A; x /∈ A|] ==> f x = arbitraryby (simp add : extensional-def )

lemma restrict-extensional [simp]: restrict f A ∈ extensional Aby (simp add : restrict-def extensional-def )

lemma compose-extensional [simp]: compose A f g ∈ extensional Aby (simp add : compose-def )

lemma extensionalityI :[| f ∈ extensional A; g ∈ extensional A;

!!x . x∈A ==> f x = g x |] ==> f = gby (force simp add : expand-fun-eq extensional-def )

lemma Inv-funcset : f ‘ A = B ==> (λx∈B . Inv A f x ) : B −> Aapply (unfold Inv-def )apply (fast intro: restrict-in-funcset someI2 )done

lemma compose-Inv-id :[| inj-on f A; f ‘ A = B |]==> compose A (λy∈B . Inv A f y) f = (λx∈A. x )

apply (simp add : compose-def )apply (rule restrict-ext , auto)apply (erule subst)apply (simp add : Inv-f-f )

5

Page 6: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

done

lemma compose-id-Inv :f ‘ A = B ==> compose B f (λy∈B . Inv A f y) = (λx∈B . x )

apply (simp add : compose-def )apply (rule restrict-ext)apply (simp add : f-Inv-f )done

end

theory funcset-def = FuncSet :

lemma funcset-definition: shows funcset A B = {f . ∀ x . x ∈ A −−> f x ∈ B }by (auto simp add : funcset-mem, unfold Pi-def , simp)

end

theory partiality = FuncSet :

lemma foo: assumes f ∈ A → B and f = id x /∈ A shows f x = xproof − from prems show ?thesis by simpqed

lemma foo2 : assumes f ∈ A → B and g ∈ B → C shows g ◦ f ∈ A → Cproof (intro funcsetI )from prems show

∧x . x ∈ A =⇒ (g ◦ f ) x ∈ C by (auto simp add : funcset-mem)

qed

lemma foo3 : assumes f ∈ A → B and g ∈ B → C and h ∈ C → D shows h◦ g ◦ f ∈ A → Dproof (intro funcsetI )from prems show

∧x . x ∈ A =⇒ (h ◦ g ◦ f ) x ∈ D by (auto simp add :

funcset-mem)qed

lemma fooo: assumes f ∈ A → B and g ∈ A → B and ∀ x . f x = g x shows f= gproof (auto simp add : expand-fun-eq) from prems show

∧x . f x = g x by simp

qed

lemma fooo2 : assumes f = (λx ::int . x + 1 − 1 ) and g = (λx . x ) shows f = gby (simp add : prems)

lemma fooo3 : assumes f ∈ {0 ::int , 1} → {0 , 1} and f (0 ) = 0 and f (1 ) = 1shows f = fby (auto simp add : expand-fun-eq)

6

Page 7: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

end

2 Groups

theory Group = FuncSet :

3 From Magmas to Groups

Definitions follow Jacobson, Basic Algebra I, Freeman, 1985; with the ex-ception of magma which, following Bourbaki, is a set together with a binary,closed operation.

3.1 Definitions

record ′a partial-object =carrier :: ′a set

record ′a semigroup = ′a partial-object +mult :: [ ′a, ′a] => ′a (infixl ⊗ı 70 )

record ′a monoid = ′a semigroup +one :: ′a (1ı)

constdefsm-inv :: [( ′a, ′m) monoid-scheme, ′a] => ′a (inv ı - [81 ] 80 )m-inv G x == (THE y . y ∈ carrier G &

mult G x y = one G & mult G y x = one G)

Units :: ( ′a, ′m) monoid-scheme => ′a setUnits G == {y . y ∈ carrier G &

(EX x : carrier G . mult G x y = one G & mult G y x = one G)}

constspow :: [( ′a, ′m) monoid-scheme, ′a, ′b::number ] => ′a (infixr ′(ˆ ′)ı 75 )

defs (overloaded)nat-pow-def : pow G a n == nat-rec (one G) (%u b. mult G b a) nint-pow-def : pow G a z ==

let p = nat-rec (one G) (%u b. mult G b a)in if neg z then m-inv G (p (nat (−z ))) else p (nat z )

locale magma = struct G +assumes m-closed [intro, simp]:

7

Page 8: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

[| x ∈ carrier G ; y ∈ carrier G |] ==> x ⊗ y ∈ carrier G

locale semigroup = magma +assumes m-assoc:

[| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>(x ⊗ y) ⊗ z = x ⊗ (y ⊗ z )

locale monoid = semigroup +assumes one-closed [intro, simp]: 1 ∈ carrier Gand l-one [simp]: x ∈ carrier G ==> 1 ⊗ x = xand r-one [simp]: x ∈ carrier G ==> x ⊗ 1 = x

lemma monoidI :assumes m-closed :

!!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==> mult G x y ∈ carrier Gand one-closed : one G ∈ carrier Gand m-assoc:

!!x y z . [| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>mult G (mult G x y) z = mult G x (mult G y z )

and l-one: !!x . x ∈ carrier G ==> mult G (one G) x = xand r-one: !!x . x ∈ carrier G ==> mult G x (one G) = x

shows monoid Gby (fast intro!: monoid .intro magma.intro semigroup-axioms.intro

semigroup.intro monoid-axioms.introintro: prems)

lemma (in monoid) Units-closed [dest ]:x ∈ Units G ==> x ∈ carrier Gby (unfold Units-def ) fast

lemma (in monoid) inv-unique:assumes eq : y ⊗ x = 1 x ⊗ y ′ = 1and G : x ∈ carrier G y ∈ carrier G y ′ ∈ carrier G

shows y = y ′

proof −from G eq have y = y ⊗ (x ⊗ y ′) by simpalso from G have ... = (y ⊗ x ) ⊗ y ′ by (simp add : m-assoc)also from G eq have ... = y ′ by simpfinally show ?thesis .

qed

lemma (in monoid) Units-one-closed [intro, simp]:1 ∈ Units Gby (unfold Units-def ) auto

lemma (in monoid) Units-inv-closed [intro, simp]:x ∈ Units G ==> inv x ∈ carrier Gapply (unfold Units-def m-inv-def , auto)apply (rule theI2 , fast)

8

Page 9: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (fast intro: inv-unique, fast)done

lemma (in monoid) Units-l-inv :x ∈ Units G ==> inv x ⊗ x = 1apply (unfold Units-def m-inv-def , auto)apply (rule theI2 , fast)apply (fast intro: inv-unique, fast)done

lemma (in monoid) Units-r-inv :x ∈ Units G ==> x ⊗ inv x = 1apply (unfold Units-def m-inv-def , auto)apply (rule theI2 , fast)apply (fast intro: inv-unique, fast)done

lemma (in monoid) Units-inv-Units [intro, simp]:x ∈ Units G ==> inv x ∈ Units G

proof −assume x : x ∈ Units Gshow inv x ∈ Units Gby (auto simp add : Units-def

intro: Units-l-inv Units-r-inv x Units-closed [OF x ])qed

lemma (in monoid) Units-l-cancel [simp]:[| x ∈ Units G ; y ∈ carrier G ; z ∈ carrier G |] ==>(x ⊗ y = x ⊗ z ) = (y = z )

proofassume eq : x ⊗ y = x ⊗ zand G : x ∈ Units G y ∈ carrier G z ∈ carrier G

then have (inv x ⊗ x ) ⊗ y = (inv x ⊗ x ) ⊗ zby (simp add : m-assoc Units-closed)

with G show y = z by (simp add : Units-l-inv)nextassume eq : y = zand G : x ∈ Units G y ∈ carrier G z ∈ carrier G

then show x ⊗ y = x ⊗ z by simpqed

lemma (in monoid) Units-inv-inv [simp]:x ∈ Units G ==> inv (inv x ) = x

proof −assume x : x ∈ Units Gthen have inv x ⊗ inv (inv x ) = inv x ⊗ xby (simp add : Units-l-inv Units-r-inv)

with x show ?thesis by (simp add : Units-closed)qed

9

Page 10: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma (in monoid) inv-inj-on-Units:inj-on (m-inv G) (Units G)

proof (rule inj-onI )fix x yassume G : x ∈ Units G y ∈ Units G and eq : inv x = inv ythen have inv (inv x ) = inv (inv y) by simpwith G show x = y by simp

qed

lemma (in monoid) Units-inv-comm:assumes inv : x ⊗ y = 1and G : x ∈ Units G y ∈ Units G

shows y ⊗ x = 1proof −from G have x ⊗ y ⊗ x = x ⊗ 1 by (auto simp add : inv Units-closed)with G show ?thesis by (simp del : r-one add : m-assoc Units-closed)

qed

Power

lemma (in monoid) nat-pow-closed [intro, simp]:x ∈ carrier G ==> x (ˆ) (n::nat) ∈ carrier Gby (induct n) (simp-all add : nat-pow-def )

lemma (in monoid) nat-pow-0 [simp]:x (ˆ) (0 ::nat) = 1by (simp add : nat-pow-def )

lemma (in monoid) nat-pow-Suc [simp]:x (ˆ) (Suc n) = x (ˆ) n ⊗ xby (simp add : nat-pow-def )

lemma (in monoid) nat-pow-one [simp]:1 (ˆ) (n::nat) = 1by (induct n) simp-all

lemma (in monoid) nat-pow-mult :x ∈ carrier G ==> x (ˆ) (n::nat) ⊗ x (ˆ) m = x (ˆ) (n + m)by (induct m) (simp-all add : m-assoc [THEN sym])

lemma (in monoid) nat-pow-pow :x ∈ carrier G ==> (x (ˆ) n) (ˆ) m = x (ˆ) (n ∗ m::nat)by (induct m) (simp, simp add : nat-pow-mult add-commute)

A group is a monoid all of whose elements are invertible.

locale group = monoid +assumes Units: carrier G <= Units G

theorem groupI :

10

Page 11: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

assumes m-closed [simp]:!!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==> mult G x y ∈ carrier G

and one-closed [simp]: one G ∈ carrier Gand m-assoc:

!!x y z . [| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>mult G (mult G x y) z = mult G x (mult G y z )

and l-one [simp]: !!x . x ∈ carrier G ==> mult G (one G) x = xand l-inv-ex : !!x . x ∈ carrier G ==> EX y : carrier G . mult G y x = one G

shows group Gproof −have l-cancel [simp]:

!!x y z . [| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>(mult G x y = mult G x z ) = (y = z )

prooffix x y zassume eq : mult G x y = mult G x zand G : x ∈ carrier G y ∈ carrier G z ∈ carrier G

with l-inv-ex obtain x-inv where xG : x-inv ∈ carrier Gand l-inv : mult G x-inv x = one G by fast

from G eq xG have mult G (mult G x-inv x ) y = mult G (mult G x-inv x ) zby (simp add : m-assoc)

with G show y = z by (simp add : l-inv)nextfix x y zassume eq : y = zand G : x ∈ carrier G y ∈ carrier G z ∈ carrier G

then show mult G x y = mult G x z by simpqedhave r-one:

!!x . x ∈ carrier G ==> mult G x (one G) = xproof −fix xassume x : x ∈ carrier Gwith l-inv-ex obtain x-inv where xG : x-inv ∈ carrier Gand l-inv : mult G x-inv x = one G by fast

from x xG have mult G x-inv (mult G x (one G)) = mult G x-inv xby (simp add : m-assoc [symmetric] l-inv)

with x xG show mult G x (one G) = x by simpqedhave inv-ex :

!!x . x ∈ carrier G ==> EX y : carrier G . mult G y x = one G &mult G x y = one G

proof −fix xassume x : x ∈ carrier Gwith l-inv-ex obtain y where y : y ∈ carrier Gand l-inv : mult G y x = one G by fast

from x y have mult G y (mult G x y) = mult G y (one G)by (simp add : m-assoc [symmetric] l-inv r-one)

11

Page 12: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

with x y have r-inv : mult G x y = one Gby simp

from x y show EX y : carrier G . mult G y x = one G &mult G x y = one Gby (fast intro: l-inv r-inv)

qedthen have carrier-subset-Units: carrier G <= Units Gby (unfold Units-def ) fast

show ?thesisby (fast intro!: group.intro magma.intro semigroup-axioms.intro

semigroup.intro monoid-axioms.intro group-axioms.introcarrier-subset-Units intro: prems r-one)

qed

lemma (in monoid) monoid-groupI :assumes l-inv-ex :

!!x . x ∈ carrier G ==> EX y : carrier G . mult G y x = one Gshows group Gby (rule groupI ) (auto intro: m-assoc l-inv-ex )

lemma (in group) Units-eq [simp]:Units G = carrier G

proofshow Units G <= carrier G by fast

nextshow carrier G <= Units G by (rule Units)

qed

lemma (in group) inv-closed [intro, simp]:x ∈ carrier G ==> inv x ∈ carrier Gusing Units-inv-closed by simp

lemma (in group) l-inv :x ∈ carrier G ==> inv x ⊗ x = 1using Units-l-inv by simp

3.2 Cancellation Laws and Basic Properties

lemma (in group) l-cancel [simp]:[| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>(x ⊗ y = x ⊗ z ) = (y = z )using Units-l-inv by simp

lemma (in group) r-inv :x ∈ carrier G ==> x ⊗ inv x = 1

proof −assume x : x ∈ carrier Gthen have inv x ⊗ (x ⊗ inv x ) = inv x ⊗ 1by (simp add : m-assoc [symmetric] l-inv)

12

Page 13: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

with x show ?thesis by (simp del : r-one)qed

lemma (in group) r-cancel [simp]:[| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>(y ⊗ x = z ⊗ x ) = (y = z )

proofassume eq : y ⊗ x = z ⊗ xand G : x ∈ carrier G y ∈ carrier G z ∈ carrier G

then have y ⊗ (x ⊗ inv x ) = z ⊗ (x ⊗ inv x )by (simp add : m-assoc [symmetric])

with G show y = z by (simp add : r-inv)nextassume eq : y = zand G : x ∈ carrier G y ∈ carrier G z ∈ carrier G

then show y ⊗ x = z ⊗ x by simpqed

lemma (in group) inv-one [simp]:inv 1 = 1

proof −have inv 1 = 1 ⊗ (inv 1) by simpmoreover have ... = 1 by (simp add : r-inv)finally show ?thesis .

qed

lemma (in group) inv-inv [simp]:x ∈ carrier G ==> inv (inv x ) = xusing Units-inv-inv by simp

lemma (in group) inv-inj :inj-on (m-inv G) (carrier G)using inv-inj-on-Units by simp

lemma (in group) inv-mult-group:[| x ∈ carrier G ; y ∈ carrier G |] ==> inv (x ⊗ y) = inv y ⊗ inv x

proof −assume G : x ∈ carrier G y ∈ carrier Gthen have inv (x ⊗ y) ⊗ (x ⊗ y) = (inv y ⊗ inv x ) ⊗ (x ⊗ y)by (simp add : m-assoc l-inv) (simp add : m-assoc [symmetric] l-inv)

with G show ?thesis by simpqed

lemma (in group) inv-comm:[| x ⊗ y = 1; x ∈ carrier G ; y ∈ carrier G |] ==> y ⊗ x = 1by (rule Units-inv-comm) auto

lemma (in group) inv-equality :[|y ⊗ x = 1; x ∈ carrier G ; y ∈ carrier G |] ==> inv x = y

13

Page 14: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (simp add : m-inv-def )apply (rule the-equality)apply (simp add : inv-comm [of y x ])apply (rule r-cancel [THEN iffD1 ], auto)done

Power

lemma (in group) int-pow-def2 :a (ˆ) (z ::int) = (if neg z then inv (a (ˆ) (nat (−z ))) else a (ˆ) (nat z ))by (simp add : int-pow-def nat-pow-def Let-def )

lemma (in group) int-pow-0 [simp]:x (ˆ) (0 ::int) = 1by (simp add : int-pow-def2 )

lemma (in group) int-pow-one [simp]:1 (ˆ) (z ::int) = 1by (simp add : int-pow-def2 )

3.3 Substructures

locale submagma = var H + struct G +assumes subset [intro, simp]: H ⊆ carrier Gand m-closed [intro, simp]: [| x ∈ H ; y ∈ H |] ==> x ⊗ y ∈ H

declare (in submagma) magma.intro [intro] semigroup.intro [intro]semigroup-axioms.intro [intro]

lemma submagma-imp-subset :submagma H G ==> H ⊆ carrier Gby (rule submagma.subset)

lemma (in submagma) subsetD [dest , simp]:x ∈ H ==> x ∈ carrier Gusing subset by blast

lemma (in submagma) magmaI [intro]:includes magma Gshows magma (G(| carrier := H |))by rule simp

lemma (in submagma) semigroup-axiomsI [intro]:includes semigroup Gshows semigroup-axioms (G(| carrier := H |))by rule (simp add : m-assoc)

lemma (in submagma) semigroupI [intro]:includes semigroup G

14

Page 15: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

shows semigroup (G(| carrier := H |))using prems by fast

locale subgroup = submagma H G +assumes one-closed [intro, simp]: 1 ∈ Hand m-inv-closed [intro, simp]: x ∈ H ==> inv x ∈ H

declare (in subgroup) group.intro [intro]

lemma (in subgroup) group-axiomsI [intro]:includes group Gshows group-axioms (G(| carrier := H |))by (rule group-axioms.intro) (auto intro: l-inv r-inv simp add : Units-def )

lemma (in subgroup) groupI [intro]:includes group Gshows group (G(| carrier := H |))by (rule groupI ) (auto intro: m-assoc l-inv)

Since H is nonempty, it contains some element x. Since it is closed underinverse, it contains inv x. Since it is closed under product, it contains x ⊗inv x = 1.

lemma (in group) one-in-subset :[| H ⊆ carrier G ; H 6= {}; ∀ a ∈ H . inv a ∈ H ; ∀ a∈H . ∀ b∈H . a ⊗ b ∈ H |]==> 1 ∈ H

by (force simp add : l-inv)

A characterization of subgroups: closed, non-empty subset.

lemma (in group) subgroupI :assumes subset : H ⊆ carrier G and non-empty : H 6= {}and inv : !!a. a ∈ H ==> inv a ∈ Hand mult : !!a b. [|a ∈ H ; b ∈ H |] ==> a ⊗ b ∈ H

shows subgroup H Gproof (rule subgroup.intro)from subset and mult show submagma H G by (rule submagma.intro)

nexthave 1 ∈ H by (rule one-in-subset) (auto simp only : prems)with inv show subgroup-axioms H Gby (intro subgroup-axioms.intro) simp-all

qed

Repeat facts of submagmas for subgroups. Necessary???

lemma (in subgroup) subset :H ⊆ carrier G..

lemma (in subgroup) m-closed :

15

Page 16: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

[| x ∈ H ; y ∈ H |] ==> x ⊗ y ∈ H..

declare magma.m-closed [simp]

declare monoid .one-closed [iff ] group.inv-closed [simp]monoid .l-one [simp] monoid .r-one [simp] group.inv-inv [simp]

lemma subgroup-nonempty :∼ subgroup {} Gby (blast dest : subgroup.one-closed)

lemma (in subgroup) finite-imp-card-positive:finite (carrier G) ==> 0 < card H

proof (rule classical)have sub: subgroup H G using prems by (rule subgroup.intro)assume fin: finite (carrier G)and zero: ∼ 0 < card H

then have finite H by (blast intro: finite-subset dest : subset)with zero sub have subgroup {} G by simpwith subgroup-nonempty show ?thesis by contradiction

qed

3.4 Direct Products

constdefsDirProdSemigroup ::

[( ′a, ′m) semigroup-scheme, ( ′b, ′n) semigroup-scheme]=> ( ′a × ′b) semigroup(infixr ×s 80 )

G ×s H == (| carrier = carrier G × carrier H ,mult = (%(xg , xh) (yg , yh). (mult G xg yg , mult H xh yh)) |)

DirProdGroup ::[( ′a, ′m) monoid-scheme, ( ′b, ′n) monoid-scheme] => ( ′a × ′b) monoid(infixr ×g 80 )

G ×g H == (| carrier = carrier (G ×s H ),mult = mult (G ×s H ),one = (one G , one H ) |)

lemma DirProdSemigroup-magma:includes magma G + magma Hshows magma (G ×s H )by (rule magma.intro) (auto simp add : DirProdSemigroup-def )

lemma DirProdSemigroup-semigroup-axioms:includes semigroup G + semigroup Hshows semigroup-axioms (G ×s H )by (rule semigroup-axioms.intro)

16

Page 17: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

(auto simp add : DirProdSemigroup-def G .m-assoc H .m-assoc)

lemma DirProdSemigroup-semigroup:includes semigroup G + semigroup Hshows semigroup (G ×s H )using premsby (fast intro: semigroup.intro

DirProdSemigroup-magma DirProdSemigroup-semigroup-axioms)

lemma DirProdGroup-magma:includes magma G + magma Hshows magma (G ×g H )by (rule magma.intro)

(auto simp add : DirProdGroup-def DirProdSemigroup-def )

lemma DirProdGroup-semigroup-axioms:includes semigroup G + semigroup Hshows semigroup-axioms (G ×g H )by (rule semigroup-axioms.intro)

(auto simp add : DirProdGroup-def DirProdSemigroup-defG .m-assoc H .m-assoc)

lemma DirProdGroup-semigroup:includes semigroup G + semigroup Hshows semigroup (G ×g H )using premsby (fast intro: semigroup.intro

DirProdGroup-magma DirProdGroup-semigroup-axioms)

lemma DirProdGroup-group:includes group G + group Hshows group (G ×g H )by (rule groupI )

(auto intro: G .m-assoc H .m-assoc G .l-inv H .l-invsimp add : DirProdGroup-def DirProdSemigroup-def )

lemma carrier-DirProdGroup [simp]:carrier (G ×g H ) = carrier G × carrier H

by (simp add : DirProdGroup-def DirProdSemigroup-def )

lemma one-DirProdGroup [simp]:one (G ×g H ) = (one G , one H )

by (simp add : DirProdGroup-def DirProdSemigroup-def )

lemma mult-DirProdGroup [simp]:mult (G ×g H ) (g , h) (g ′, h ′) = (mult G g g ′, mult H h h ′)

by (simp add : DirProdGroup-def DirProdSemigroup-def )

17

Page 18: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma inv-DirProdGroup [simp]:includes group G + group Hassumes g : g ∈ carrier G

and h: h ∈ carrier Hshows m-inv (G ×g H ) (g , h) = (m-inv G g , m-inv H h)apply (rule group.inv-equality [OF DirProdGroup-group])apply (simp-all add : prems group-def group.l-inv)done

3.5 Homomorphisms

constdefshom :: [( ′a, ′c) semigroup-scheme, ( ′b, ′d) semigroup-scheme]

=> ( ′a => ′b)sethom G H =={h. h ∈ carrier G −> carrier H &

(∀ x ∈ carrier G . ∀ y ∈ carrier G . h (mult G x y) = mult H (h x ) (h y))}

lemma (in semigroup) hom:includes semigroup Gshows semigroup (| carrier = hom G G , mult = op o |)

proof (rule semigroup.intro)show magma (| carrier = hom G G , mult = op o |)by (rule magma.intro) (simp add : Pi-def hom-def )

nextshow semigroup-axioms (| carrier = hom G G , mult = op o |)by (rule semigroup-axioms.intro) (simp add : o-assoc)

qed

lemma hom-mult :[| h ∈ hom G H ; x ∈ carrier G ; y ∈ carrier G |]==> h (mult G x y) = mult H (h x ) (h y)by (simp add : hom-def )

lemma hom-closed :[| h ∈ hom G H ; x ∈ carrier G |] ==> h x ∈ carrier Hby (auto simp add : hom-def funcset-mem)

lemma compose-hom:[|group G ; h ∈ hom G G ; h ′ ∈ hom G G ; h ′ ∈ carrier G −> carrier G |]==> compose (carrier G) h h ′ ∈ hom G G

apply (simp (no-asm-simp) add : hom-def )apply (intro conjI )apply (force simp add : funcset-compose hom-def )apply (simp add : compose-def group.axioms hom-mult funcset-mem)done

locale group-hom = group G + group H + var h +

18

Page 19: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

assumes homh: h ∈ hom G Hnotes hom-mult [simp] = hom-mult [OF homh]and hom-closed [simp] = hom-closed [OF homh]

lemma (in group-hom) one-closed [simp]:h 1 ∈ carrier Hby simp

lemma (in group-hom) hom-one [simp]:h 1 = 12

proof −have h 1 ⊗2 12 = h 1 ⊗2 h 1by (simp add : hom-mult [symmetric] del : hom-mult)

then show ?thesis by (simp del : r-one)qed

lemma (in group-hom) inv-closed [simp]:x ∈ carrier G ==> h (inv x ) ∈ carrier Hby simp

lemma (in group-hom) hom-inv [simp]:x ∈ carrier G ==> h (inv x ) = inv2 (h x )

proof −assume x : x ∈ carrier Gthen have h x ⊗2 h (inv x ) = 12

by (simp add : hom-mult [symmetric] G .r-inv del : hom-mult)also from x have ... = h x ⊗2 inv2 (h x )by (simp add : hom-mult [symmetric] H .r-inv del : hom-mult)

finally have h x ⊗2 h (inv x ) = h x ⊗2 inv2 (h x ) .with x show ?thesis by simp

qed

3.6 Commutative Structures

Naming convention: multiplicative structures that are commutative arecalled commutative, additive structures are called Abelian.

3.7 Definition

locale comm-semigroup = semigroup +assumes m-comm: [| x ∈ carrier G ; y ∈ carrier G |] ==> x ⊗ y = y ⊗ x

lemma (in comm-semigroup) m-lcomm:[| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>x ⊗ (y ⊗ z ) = y ⊗ (x ⊗ z )

proof −assume xyz : x ∈ carrier G y ∈ carrier G z ∈ carrier Gfrom xyz have x ⊗ (y ⊗ z ) = (x ⊗ y) ⊗ z by (simp add : m-assoc)also from xyz have ... = (y ⊗ x ) ⊗ z by (simp add : m-comm)

19

Page 20: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

also from xyz have ... = y ⊗ (x ⊗ z ) by (simp add : m-assoc)finally show ?thesis .

qed

lemmas (in comm-semigroup) m-ac = m-assoc m-comm m-lcomm

locale comm-monoid = comm-semigroup + monoid

lemma comm-monoidI :assumes m-closed :

!!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==> mult G x y ∈ carrier Gand one-closed : one G ∈ carrier Gand m-assoc:

!!x y z . [| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>mult G (mult G x y) z = mult G x (mult G y z )

and l-one: !!x . x ∈ carrier G ==> mult G (one G) x = xand m-comm:

!!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==> mult G x y = mult G y xshows comm-monoid Gusing l-oneby (auto intro!: comm-monoid .intro magma.intro semigroup-axioms.intro

comm-semigroup-axioms.intro monoid-axioms.introintro: prems simp: m-closed one-closed m-comm)

lemma (in monoid) monoid-comm-monoidI :assumes m-comm:

!!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==> mult G x y = mult G y xshows comm-monoid Gby (rule comm-monoidI ) (auto intro: m-assoc m-comm)

lemma (in comm-monoid) nat-pow-distr :[| x ∈ carrier G ; y ∈ carrier G |] ==>(x ⊗ y) (ˆ) (n::nat) = x (ˆ) n ⊗ y (ˆ) nby (induct n) (simp, simp add : m-ac)

locale comm-group = comm-monoid + group

lemma (in group) group-comm-groupI :assumes m-comm: !!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==>

mult G x y = mult G y xshows comm-group Gby (fast intro: comm-group.intro comm-semigroup-axioms.intro

group.axioms prems)

lemma comm-groupI :assumes m-closed :

!!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==> mult G x y ∈ carrier Gand one-closed : one G ∈ carrier G

20

Page 21: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

and m-assoc:!!x y z . [| x ∈ carrier G ; y ∈ carrier G ; z ∈ carrier G |] ==>mult G (mult G x y) z = mult G x (mult G y z )

and m-comm:!!x y . [| x ∈ carrier G ; y ∈ carrier G |] ==> mult G x y = mult G y x

and l-one: !!x . x ∈ carrier G ==> mult G (one G) x = xand l-inv-ex : !!x . x ∈ carrier G ==> EX y : carrier G . mult G y x = one G

shows comm-group Gby (fast intro: group.group-comm-groupI groupI prems)

lemma (in comm-group) inv-mult :[| x ∈ carrier G ; y ∈ carrier G |] ==> inv (x ⊗ y) = inv x ⊗ inv yby (simp add : m-ac inv-mult-group)

end

4 Product Operator

theory FiniteProduct = Group:

constsfoldSetD :: [ ′a set , ′b => ′a => ′a, ′a] => ( ′b set ∗ ′a) set

inductive foldSetD D f eintros

emptyI [intro]: e : D ==> ({}, e) : foldSetD D f einsertI [intro]: [| x ∼: A; f x y : D ; (A, y) : foldSetD D f e |] ==>

(insert x A, f x y) : foldSetD D f e

inductive-cases empty-foldSetDE [elim!]: ({}, x ) : foldSetD D f e

constdefsfoldD :: [ ′a set , ′b => ′a => ′a, ′a, ′b set ] => ′afoldD D f e A == THE x . (A, x ) : foldSetD D f e

lemma foldSetD-closed :[| (A, z ) : foldSetD D f e ; e : D ; !!x y . [| x : A; y : D |] ==> f x y : D

|] ==> z : Dby (erule foldSetD .elims) auto

lemma Diff1-foldSetD :[| (A − {x}, y) : foldSetD D f e; x : A; f x y : D |] ==>(A, f x y) : foldSetD D f eapply (erule insert-Diff [THEN subst ], rule foldSetD .intros)apply auto

done

21

Page 22: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma foldSetD-imp-finite [simp]: (A, x ) : foldSetD D f e ==> finite Aby (induct set : foldSetD) auto

lemma finite-imp-foldSetD :[| finite A; e : D ; !!x y . [| x : A; y : D |] ==> f x y : D |] ==>EX x . (A, x ) : foldSetD D f e

proof (induct set : Finites)case empty then show ?case by auto

nextcase (insert F x )then obtain y where y : (F , y) : foldSetD D f e by autowith insert have y : D by (auto dest : foldSetD-closed)with y and insert have (insert x F , f x y) : foldSetD D f eby (intro foldSetD .intros) auto

then show ?case ..qed

4.1 Left-commutative operations

locale LCD =fixes B :: ′b setand D :: ′a setand f :: ′b => ′a => ′a (infixl · 70 )assumes left-commute:

[| x : B ; y : B ; z : D |] ==> x · (y · z ) = y · (x · z )and f-closed [simp, intro!]: !!x y . [| x : B ; y : D |] ==> f x y : D

lemma (in LCD) foldSetD-closed [dest ]:(A, z ) : foldSetD D f e ==> z : Dby (erule foldSetD .elims) auto

lemma (in LCD) Diff1-foldSetD :[| (A − {x}, y) : foldSetD D f e; x : A; A <= B |] ==>(A, f x y) : foldSetD D f eapply (subgoal-tac x : B)prefer 2 apply fastapply (erule insert-Diff [THEN subst ], rule foldSetD .intros)apply auto

done

lemma (in LCD) foldSetD-imp-finite [simp]:(A, x ) : foldSetD D f e ==> finite Aby (induct set : foldSetD) auto

lemma (in LCD) finite-imp-foldSetD :[| finite A; A <= B ; e : D |] ==> EX x . (A, x ) : foldSetD D f e

proof (induct set : Finites)case empty then show ?case by auto

22

Page 23: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

nextcase (insert F x )then obtain y where y : (F , y) : foldSetD D f e by autowith insert have y : D by autowith y and insert have (insert x F , f x y) : foldSetD D f eby (intro foldSetD .intros) auto

then show ?case ..qed

lemma (in LCD) foldSetD-determ-aux :e : D ==> ALL A x . A <= B & card A < n −−> (A, x ) : foldSetD D f e −−>

(ALL y . (A, y) : foldSetD D f e −−> y = x )apply (induct n)apply (auto simp add : less-Suc-eq)apply (erule foldSetD .cases)apply blastapply (erule foldSetD .cases)apply blastapply clarify

force simplification of card A < card (insert ...).

apply (erule rev-mp)apply (simp add : less-Suc-eq-le)apply (rule impI )apply (rename-tac Aa xa ya Ab xb yb, case-tac xa = xb)apply (subgoal-tac Aa = Ab)prefer 2 apply (blast elim!: equalityE )apply blast

case xa /∈ xb.

apply (subgoal-tac Aa − {xb} = Ab − {xa} & xb : Aa & xa : Ab)prefer 2 apply (blast elim!: equalityE )apply clarifyapply (subgoal-tac Aa = insert xb Ab − {xa})prefer 2 apply blastapply (subgoal-tac card Aa <= card Ab)prefer 2apply (rule Suc-le-mono [THEN subst ])apply (simp add : card-Suc-Diff1 )apply (rule-tac A1 = Aa − {xb} in finite-imp-foldSetD [THEN exE ])

apply (blast intro: foldSetD-imp-finite finite-Diff )apply bestapply assumptionapply (frule (1 ) Diff1-foldSetD)apply bestapply (subgoal-tac ya = f xb x )prefer 2apply (subgoal-tac Aa <= B)prefer 2 apply best

23

Page 24: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (blast del : equalityCE )apply (subgoal-tac (Ab − {xa}, x ) : foldSetD D f e)prefer 2 apply simpapply (subgoal-tac yb = f xa x )prefer 2apply (blast del : equalityCE dest : Diff1-foldSetD)apply (simp (no-asm-simp))apply (rule left-commute)apply assumptionapply bestapply bestdone

lemma (in LCD) foldSetD-determ:[| (A, x ) : foldSetD D f e; (A, y) : foldSetD D f e; e : D ; A <= B |]==> y = xby (blast intro: foldSetD-determ-aux [rule-format ])

lemma (in LCD) foldD-equality :[| (A, y) : foldSetD D f e; e : D ; A <= B |] ==> foldD D f e A = yby (unfold foldD-def ) (blast intro: foldSetD-determ)

lemma foldD-empty [simp]:e : D ==> foldD D f e {} = eby (unfold foldD-def ) blast

lemma (in LCD) foldD-insert-aux :[| x ∼: A; x : B ; e : D ; A <= B |] ==>

((insert x A, v) : foldSetD D f e) =(EX y . (A, y) : foldSetD D f e & v = f x y)

apply autoapply (rule-tac A1 = A in finite-imp-foldSetD [THEN exE ])

apply (fastsimp dest : foldSetD-imp-finite)apply assumptionapply assumptionapply (blast intro: foldSetD-determ)done

lemma (in LCD) foldD-insert :[| finite A; x ∼: A; x : B ; e : D ; A <= B |] ==>foldD D f e (insert x A) = f x (foldD D f e A)

apply (unfold foldD-def )apply (simp add : foldD-insert-aux )apply (rule the-equality)apply (auto intro: finite-imp-foldSetD

cong add : conj-cong simp add : foldD-def [symmetric] foldD-equality)done

lemma (in LCD) foldD-closed [simp]:

24

Page 25: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

[| finite A; e : D ; A <= B |] ==> foldD D f e A : Dproof (induct set : Finites)case empty then show ?case by (simp add : foldD-empty)

nextcase insert then show ?case by (simp add : foldD-insert)

qed

lemma (in LCD) foldD-commute:[| finite A; x : B ; e : D ; A <= B |] ==>f x (foldD D f e A) = foldD D f (f x e) Aapply (induct set : Finites)apply simpapply (auto simp add : left-commute foldD-insert)done

lemma Int-mono2 :[| A <= C ; B <= C |] ==> A Int B <= Cby blast

lemma (in LCD) foldD-nest-Un-Int :[| finite A; finite C ; e : D ; A <= B ; C <= B |] ==>foldD D f (foldD D f e C ) A = foldD D f (foldD D f e (A Int C )) (A Un C )apply (induct set : Finites)apply simpapply (simp add : foldD-insert foldD-commute Int-insert-left insert-absorb

Int-mono2 Un-subset-iff )done

lemma (in LCD) foldD-nest-Un-disjoint :[| finite A; finite B ; A Int B = {}; e : D ; A <= B ; C <= B |]

==> foldD D f e (A Un B) = foldD D f (foldD D f e B) Aby (simp add : foldD-nest-Un-Int)

— Delete rules to do with foldSetD relation.

declare foldSetD-imp-finite [simp del ]empty-foldSetDE [rule del ]foldSetD .intros [rule del ]

declare (in LCD)foldSetD-closed [rule del ]

4.2 Commutative monoids

We enter a more restrictive context, with f :: ′a => ′a => ′a instead of ′b=> ′a => ′a.

locale ACeD =fixes D :: ′a setand f :: ′a => ′a => ′a (infixl · 70 )and e :: ′a

25

Page 26: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

assumes ident [simp]: x : D ==> x · e = xand commute: [| x : D ; y : D |] ==> x · y = y · xand assoc: [| x : D ; y : D ; z : D |] ==> (x · y) · z = x · (y · z )and e-closed [simp]: e : Dand f-closed [simp]: [| x : D ; y : D |] ==> x · y : D

lemma (in ACeD) left-commute:[| x : D ; y : D ; z : D |] ==> x · (y · z ) = y · (x · z )

proof −assume D : x : D y : D z : Dthen have x · (y · z ) = (y · z ) · x by (simp add : commute)also from D have ... = y · (z · x ) by (simp add : assoc)also from D have z · x = x · z by (simp add : commute)finally show ?thesis .

qed

lemmas (in ACeD) AC = assoc commute left-commute

lemma (in ACeD) left-ident [simp]: x : D ==> e · x = xproof −assume D : x : Dhave x · e = x by (rule ident)with D show ?thesis by (simp add : commute)

qed

lemma (in ACeD) foldD-Un-Int :[| finite A; finite B ; A <= D ; B <= D |] ==>

foldD D f e A · foldD D f e B =foldD D f e (A Un B) · foldD D f e (A Int B)

apply (induct set : Finites)apply (simp add : left-commute LCD .foldD-closed [OF LCD .intro [of D ]])apply (simp add : AC insert-absorb Int-insert-left

LCD .foldD-insert [OF LCD .intro [of D ]]LCD .foldD-closed [OF LCD .intro [of D ]]Int-mono2 Un-subset-iff )

done

lemma (in ACeD) foldD-Un-disjoint :[| finite A; finite B ; A Int B = {}; A <= D ; B <= D |] ==>

foldD D f e (A Un B) = foldD D f e A · foldD D f e Bby (simp add : foldD-Un-Int

left-commute LCD .foldD-closed [OF LCD .intro [of D ]] Un-subset-iff )

4.3 Products over Finite Sets

constdefsfinprod :: [( ′b, ′m) monoid-scheme, ′a => ′b, ′a set ] => ′bfinprod G f A == if finite A

then foldD (carrier G) (mult G o f ) (one G) A

26

Page 27: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

else arbitrary

ML-setup {∗simpset-ref () := simpset() setsubgoaler asm-full-simp-tac;

∗}

lemma (in comm-monoid) finprod-empty [simp]:finprod G f {} = 1by (simp add : finprod-def )

ML-setup {∗simpset-ref () := simpset() setsubgoaler asm-simp-tac;

∗}

declare funcsetI [intro]funcset-mem [dest ]

lemma (in comm-monoid) finprod-insert [simp]:[| finite F ; a /∈ F ; f ∈ F −> carrier G ; f a ∈ carrier G |] ==>finprod G f (insert a F ) = f a ⊗ finprod G f Fapply (rule trans)apply (simp add : finprod-def )apply (rule trans)apply (rule LCD .foldD-insert [OF LCD .intro [of insert a F ]])

apply simpapply (rule m-lcomm)apply fastapply fastapply assumptionapply (fastsimp intro: m-closed)apply simp+

apply fastapply (auto simp add : finprod-def )done

lemma (in comm-monoid) finprod-one [simp]:finite A ==> finprod G (%i . 1) A = 1

proof (induct set : Finites)case empty show ?case by simp

nextcase (insert A a)have (%i . 1) ∈ A −> carrier G by autowith insert show ?case by simp

qed

lemma (in comm-monoid) finprod-closed [simp]:fixes A

27

Page 28: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

assumes fin: finite A and f : f ∈ A −> carrier Gshows finprod G f A ∈ carrier G

using fin fproof inductcase empty show ?case by simp

nextcase (insert A a)then have a: f a ∈ carrier G by fastfrom insert have A: f ∈ A −> carrier G by fastfrom insert A a show ?case by simp

qed

lemma funcset-Int-left [simp, intro]:[| f ∈ A −> C ; f ∈ B −> C |] ==> f ∈ A Int B −> Cby fast

lemma funcset-Un-left [iff ]:(f ∈ A Un B −> C ) = (f ∈ A −> C & f ∈ B −> C )by fast

lemma (in comm-monoid) finprod-Un-Int :[| finite A; finite B ; g ∈ A −> carrier G ; g ∈ B −> carrier G |] ==>

finprod G g (A Un B) ⊗ finprod G g (A Int B) =finprod G g A ⊗ finprod G g B

— The reversed orientation looks more natural, but LOOPS as a simprule!proof (induct set : Finites)case empty then show ?case by (simp add : finprod-closed)

nextcase (insert A a)then have a: g a ∈ carrier G by fastfrom insert have A: g ∈ A −> carrier G by fastfrom insert A a show ?caseby (simp add : m-ac Int-insert-left insert-absorb finprod-closed

Int-mono2 Un-subset-iff )qed

lemma (in comm-monoid) finprod-Un-disjoint :[| finite A; finite B ; A Int B = {};

g ∈ A −> carrier G ; g ∈ B −> carrier G |]==> finprod G g (A Un B) = finprod G g A ⊗ finprod G g Bapply (subst finprod-Un-Int [symmetric])

apply (auto simp add : finprod-closed)done

lemma (in comm-monoid) finprod-multf :[| finite A; f ∈ A −> carrier G ; g ∈ A −> carrier G |] ==>finprod G (%x . f x ⊗ g x ) A = (finprod G f A ⊗ finprod G g A)

proof (induct set : Finites)case empty show ?case by simp

28

Page 29: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

nextcase (insert A a) thenhave fA: f : A −> carrier G by fastfrom insert have fa: f a : carrier G by fastfrom insert have gA: g : A −> carrier G by fastfrom insert have ga: g a : carrier G by fastfrom insert have fgA: (%x . f x ⊗ g x ) : A −> carrier Gby (simp add : Pi-def )

show ?caseby (simp add : insert fA fa gA ga fgA m-ac Int-insert-left insert-absorb

Int-mono2 Un-subset-iff )qed

lemma (in comm-monoid) finprod-cong ′:[| A = B ; g : B −> carrier G ;

!!i . i : B ==> f i = g i |] ==> finprod G f A = finprod G g Bproof −assume prems: A = B g : B −> carrier G

!!i . i : B ==> f i = g ishow ?thesisproof (cases finite B)case Truethen have !!A. [| A = B ; g : B −> carrier G ;

!!i . i : B ==> f i = g i |] ==> finprod G f A = finprod G g Bproof inductcase empty thus ?case by simp

nextcase (insert B x )then have finprod G f A = finprod G f (insert x B) by simpalso from insert have ... = f x ⊗ finprod G f Bproof (intro finprod-insert)show finite B .

nextshow x ∼: B .

nextassume x ∼: B !!i . i ∈ insert x B =⇒ f i = g i

g ∈ insert x B → carrier Gthus f : B −> carrier G by fastsimp

nextassume x ∼: B !!i . i ∈ insert x B =⇒ f i = g i

g ∈ insert x B → carrier Gthus f x ∈ carrier G by fastsimp

qedalso from insert have ... = g x ⊗ finprod G g B by fastsimpalso from insert have ... = finprod G g (insert x B)by (intro finprod-insert [THEN sym]) autofinally show ?case .

qedwith prems show ?thesis by simp

29

Page 30: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

nextcase False with prems show ?thesis by (simp add : finprod-def )

qedqed

lemma (in comm-monoid) finprod-cong :[| A = B ; f : B −> carrier G = True;

!!i . i : B ==> f i = g i |] ==> finprod G f A = finprod G g B

by (rule finprod-cong ′) force+

Usually, if this rule causes a failed congruence proof error, the reason is thatthe premise g : B −> carrier G cannot be shown. Adding Pi-def to thesimpset is often useful. For this reason, comm-monoid .finprod-cong is notadded to the simpset by default.

declare funcsetI [rule del ]funcset-mem [rule del ]

lemma (in comm-monoid) finprod-0 [simp]:f : {0 ::nat} −> carrier G ==> finprod G f {..0} = f 0

by (simp add : Pi-def )

lemma (in comm-monoid) finprod-Suc [simp]:f : {..Suc n} −> carrier G ==>finprod G f {..Suc n} = (f (Suc n) ⊗ finprod G f {..n})

by (simp add : Pi-def atMost-Suc)

lemma (in comm-monoid) finprod-Suc2 :f : {..Suc n} −> carrier G ==>finprod G f {..Suc n} = (finprod G (%i . f (Suc i)) {..n} ⊗ f 0 )

proof (induct n)case 0 thus ?case by (simp add : Pi-def )

nextcase Suc thus ?case by (simp add : m-assoc Pi-def )

qed

lemma (in comm-monoid) finprod-mult [simp]:[| f : {..n} −> carrier G ; g : {..n} −> carrier G |] ==>

finprod G (%i . f i ⊗ g i) {..n::nat} =finprod G f {..n} ⊗ finprod G g {..n}

by (induct n) (simp-all add : m-ac Pi-def )

end

30

Page 31: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

5 Extracting code from tuples definition and op-erations

theory composition-with-lemmas = Group:

5.1 Definition of complections

constdefscomplection :: [( ′a, ′c) monoid-scheme, ( ′b, ′d) monoid-scheme, ( ′a => ′b)] =>

( ′a => ′b)complection G G ′ f == (%x . if x ∈ carrier G then f x else one G ′)

constdefscomplection-fun :: [( ′a, ′c) monoid-scheme, ( ′b, ′d) monoid-scheme] => ( ′a =>

′b) setcomplection-fun G G ′ == {f . ∃ g . f = complection G G ′ g}

constdefshom-complection :: [( ′a, ′c) monoid-scheme, ( ′b, ′d) monoid-scheme] => ( ′a =>

′b)sethom-complection G G ′ =={h. h ∈ complection-fun G G ′ & h ∈ hom G G ′}

lemma hom-complection-closed :[| h ∈ hom-complection G G ′; x ∈ carrier G |] ==> h x ∈ carrier G ′

by (auto simp add : hom-complection-def hom-def funcset-mem)

lemma hom-complection-closed2 :[| h ∈ hom-complection G G ′; x /∈ carrier G |] ==> h x = one G ′

by (auto simp add : hom-complection-def complection-fun-def complection-def )

5.2 Definition of tuples

record ′a diff-group-type = ′a monoid +diff :: ′a ⇒ ′a (differ ı 81 )

record ( ′a, ′b) group-mrp-type =src :: ′a diff-group-typetrg :: ′b diff-group-typemorph :: ′a ⇒ ′bsrc-comm-gr :: ′a diff-group-typetrg-comm-gr :: ′b diff-group-type

constdefsgroup-mrp-comp :: [ ( ′b, ′c) group-mrp-type, ( ′a, ′b) group-mrp-type] => ( ′a, ′c)

group-mrp-typegroup-mrp-comp g f == (| src = src f , trg = trg g , morph = (morph g) ◦

(morph f ),

31

Page 32: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

src-comm-gr = src-comm-gr f , trg-comm-gr = trg-comm-gr g |)

syntax (xsymbols)group-mrp-comp :: [ ( ′b, ′c)group-mrp-type, ( ′a, ′b)group-mrp-type] => ( ′a, ′c)

group-mrp-type(infixl ◦ 55 )

constdefsdiff-group :: ′a diff-group-type => booldiff-group CC == magma CC ∧ semigroup-axioms CC ∧ comm-semigroup-axioms

CC ∧ monoid-axioms CC ∧ group-axioms CC ∧diff CC ∈ hom-complection CC CC ∧(∀ x . x ∈ carrier CC −→ diff CC (diff CC x ) = one CC )

constdefssub-diff-group :: [ ′a diff-group-type, ′a diff-group-type ] => boolsub-diff-group G G ′ == diff-group G ∧ diff-group G ′ ∧ (carrier G) ⊆ carrier G ′

∧ G = (| carrier = carrier G , mult = mult G ′, one = one G ′, diff = diff G ′|)

constdefsgroup-mrp :: ( ′a, ′b) group-mrp-type => boolgroup-mrp MRP == diff-group (src MRP) ∧ diff-group (src-comm-gr MRP) ∧

sub-diff-group (src MRP) (src-comm-gr MRP) ∧diff-group (trg MRP) ∧ diff-group (trg-comm-gr MRP) ∧ (morph

MRP) ‘ (carrier (src MRP)) ⊆ carrier (trg MRP)∧ morph MRP ∈ hom-complection (src-comm-gr MRP) (trg-comm-gr

MRP)

constdefsdiff-group-mrp :: ( ′a, ′b) group-mrp-type => booldiff-group-mrp diff-MRP == group-mrp diff-MRP ∧

(∀ x ∈ carrier (src diff-MRP). (morph diff-MRP) (diff (src diff-MRP)x ) = diff (trg diff-MRP) (morph diff-MRP x ))

5.3 Lemmas about composition of tuples

lemma group-mrp-composition:assumes A1 : group-mrp Aand B1 : group-mrp Band C1 : trg-comm-gr A = src-comm-gr Band D1 : trg A = src Bshows group-mrp (B ◦ A)

proof (unfold group-mrp-comp-def group-mrp-def , intro conjI , simp-all)from prems show diff-group (src A) by (simp add : group-mrp-def )from prems show diff-group (src-comm-gr A) by (simp add : group-mrp-def )from prems show sub-diff-group (src A) (src-comm-gr A) by (simp add : group-mrp-def )from prems show diff-group (trg B) by (simp add : group-mrp-def )from prems show diff-group (trg-comm-gr B) by (simp add : group-mrp-def )

32

Page 33: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from prems show (morph B o morph A) ‘ (carrier (src A)) ⊆ carrier (trg B)by (unfold group-mrp-def image-def , auto)

from prems show morph B ◦ morph A ∈ hom-complection (src-comm-gr A)(trg-comm-gr B)proof (unfold hom-complection-def group-mrp-def

complection-fun-def complection-def , simp, intro conjI )assume group-mrp A and group-mrp B and trg-comm-gr A = src-comm-gr Bthen show EX g . morph B o morph A = (%x . if x : carrier (src-comm-gr A)

then g x else one (trg-comm-gr B))proof (intro exI [of - morph B ◦ morph A], auto simp add : expand-fun-eq

group-mrp-def )fix x assume x /∈ carrier (src-comm-gr A)

and morph A : hom-complection (src-comm-gr A) (src-comm-gr B)and morph B : hom-complection (src-comm-gr B) (trg-comm-gr B)show morph B (morph A x ) = one (trg-comm-gr B)proof−

from prems have morph A x = one (src-comm-gr B) by (intro hom-complection-closed2 )then have morph B (morph A x ) = morph B (one (src-comm-gr B)) by simp

also from prems have ... = one (trg-comm-gr B)by (intro group-hom.hom-one [of - - morph B ],

unfold group-hom-def group-hom-axioms-def group-mrp-def diff-group-defhom-complection-def , simp)

finally show ?thesis by simpqed

qednext

from prems show morph B ◦ morph A ∈ hom (src-comm-gr A) (trg-comm-grB)

by (unfold hom-def , simp, intro conjI , unfold Pi-def group-mrp-def hom-complection-defhom-def , simp)

(auto simp add : magma.m-closed)qed

qed

lemma diff-group-mrp-composition:assumes A1 : diff-group-mrp Aand B1 : diff-group-mrp Band C1 : trg-comm-gr A = src-comm-gr Band D1 : trg A = src Bshows group-mrp (B ◦ A)

proof (rule group-mrp-composition)from prems show group-mrp A by (simp add : diff-group-mrp-def )from prems show group-mrp B by (simp add : diff-group-mrp-def )from prems show trg-comm-gr A = src-comm-gr B by simpfrom prems show trg A = src B by simp

qed

33

Page 34: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

5.4 Code extraction

types-codeset ((- −> bool))

consts-codeCollect ((-))op : ((- |> -))

generate-code (/home/jesules/Desktop/Composition.ML)comp1 = compcomp2 = group-mrp-comp

end

6 The Greatest Common Divisor and Euclid’s al-gorithm Christophe Tabacznyj and Lawrence CPaulson

theory Primes = Main:

See [?].

constsgcd :: nat × nat => nat — Euclid’s algorithm

recdef gcd measure ((λ(m, n). n) :: nat × nat => nat)gcd (m, n) = (if n = 0 then m else gcd (n, m mod n))

constdefsis-gcd :: nat => nat => nat => bool — gcd as a relationis-gcd p m n == p dvd m ∧ p dvd n ∧

(∀ d . d dvd m ∧ d dvd n −−> d dvd p)

coprime :: nat => nat => boolcoprime m n == gcd (m, n) = 1

prime :: nat setprime == {p. 1 < p ∧ (∀m. m dvd p −−> m = 1 ∨ m = p)}

lemma gcd-induct :(!!m. P m 0 ) ==>

(!!m n. 0 < n ==> P n (m mod n) ==> P m n)==> P (m::nat) (n::nat)apply (induct m n rule: gcd .induct)apply (case-tac n = 0 )apply simp-all

34

Page 35: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

done

lemma gcd-0 [simp]: gcd (m, 0 ) = mapply simpdone

lemma gcd-non-0 : 0 < n ==> gcd (m, n) = gcd (n, m mod n)apply simpdone

declare gcd .simps [simp del ]

lemma gcd-1 [simp]: gcd (m, Suc 0 ) = 1apply (simp add : gcd-non-0 )done

gcd (m, n) divides m and n. The conjunctions don’t seem provable sepa-rately.

lemma gcd-dvd1 [iff ]: gcd (m, n) dvd mand gcd-dvd2 [iff ]: gcd (m, n) dvd napply (induct m n rule: gcd-induct)apply (simp-all add : gcd-non-0 )apply (blast dest : dvd-mod-imp-dvd)done

Maximality: for all m, n, k naturals, if k divides m and k divides n then kdivides gcd (m, n).

lemma gcd-greatest : k dvd m ==> k dvd n ==> k dvd gcd (m, n)apply (induct m n rule: gcd-induct)apply (simp-all add : gcd-non-0 dvd-mod)done

lemma gcd-greatest-iff [iff ]: (k dvd gcd (m, n)) = (k dvd m ∧ k dvd n)apply (blast intro!: gcd-greatest intro: dvd-trans)done

lemma gcd-zero: (gcd (m, n) = 0 ) = (m = 0 ∧ n = 0 )by (simp only : dvd-0-left-iff [THEN sym] gcd-greatest-iff )

Function gcd yields the Greatest Common Divisor.

lemma is-gcd : is-gcd (gcd (m, n)) m napply (simp add : is-gcd-def gcd-greatest)done

Uniqueness of GCDs.

35

Page 36: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma is-gcd-unique: is-gcd m a b ==> is-gcd n a b ==> m = napply (simp add : is-gcd-def )apply (blast intro: dvd-anti-sym)done

lemma is-gcd-dvd : is-gcd m a b ==> k dvd a ==> k dvd b ==> k dvd mapply (auto simp add : is-gcd-def )done

Commutativity

lemma is-gcd-commute: is-gcd k m n = is-gcd k n mapply (auto simp add : is-gcd-def )done

lemma gcd-commute: gcd (m, n) = gcd (n, m)apply (rule is-gcd-unique)apply (rule is-gcd)apply (subst is-gcd-commute)apply (simp add : is-gcd)done

lemma gcd-assoc: gcd (gcd (k , m), n) = gcd (k , gcd (m, n))apply (rule is-gcd-unique)apply (rule is-gcd)apply (simp add : is-gcd-def )apply (blast intro: dvd-trans)done

lemma gcd-0-left [simp]: gcd (0 , m) = mapply (simp add : gcd-commute [of 0 ])done

lemma gcd-1-left [simp]: gcd (Suc 0 , m) = 1apply (simp add : gcd-commute [of Suc 0 ])done

Multiplication laws

lemma gcd-mult-distrib2 : k ∗ gcd (m, n) = gcd (k ∗ m, k ∗ n)— [?, page 27]

apply (induct m n rule: gcd-induct)apply simpapply (case-tac k = 0 )apply (simp-all add : mod-geq gcd-non-0 mod-mult-distrib2 )done

lemma gcd-mult [simp]: gcd (k , k ∗ n) = kapply (rule gcd-mult-distrib2 [of k 1 n, simplified , symmetric])done

36

Page 37: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma gcd-self [simp]: gcd (k , k) = kapply (rule gcd-mult [of k 1 , simplified ])done

lemma relprime-dvd-mult : gcd (k , n) = 1 ==> k dvd m ∗ n ==> k dvd mapply (insert gcd-mult-distrib2 [of m k n])apply simpapply (erule-tac t = m in ssubst)apply simpdone

lemma relprime-dvd-mult-iff : gcd (k , n) = 1 ==> (k dvd m ∗ n) = (k dvd m)apply (blast intro: relprime-dvd-mult dvd-trans)done

lemma prime-imp-relprime: p ∈ prime ==> ¬ p dvd n ==> gcd (p, n) = 1apply (auto simp add : prime-def )apply (drule-tac x = gcd (p, n) in spec)apply autoapply (insert gcd-dvd2 [of p n])apply simpdone

lemma two-is-prime: 2 ∈ primeapply (auto simp add : prime-def )apply (case-tac m)apply (auto dest !: dvd-imp-le)done

This theorem leads immediately to a proof of the uniqueness of factorization.If p divides a product of primes then it is one of those primes.

lemma prime-dvd-mult : p ∈ prime ==> p dvd m ∗ n ==> p dvd m ∨ p dvd nby (blast intro: relprime-dvd-mult prime-imp-relprime)

lemma prime-dvd-square: p ∈ prime ==> p dvd mˆSuc (Suc 0 ) ==> p dvd mby (auto dest : prime-dvd-mult)

lemma prime-dvd-power-two: p ∈ prime ==> p dvd m ==> p dvd mby (rule prime-dvd-square) (simp-all add : power2-eq-square)

Addition laws

lemma gcd-add1 [simp]: gcd (m + n, n) = gcd (m, n)apply (case-tac n = 0 )apply (simp-all add : gcd-non-0 )done

lemma gcd-add2 [simp]: gcd (m, m + n) = gcd (m, n)

37

Page 38: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (rule gcd-commute [THEN trans])apply (subst add-commute)apply (simp add : gcd-add1 )apply (rule gcd-commute)done

lemma gcd-add2 ′ [simp]: gcd (m, n + m) = gcd (m, n)apply (subst add-commute)apply (rule gcd-add2 )done

lemma gcd-add-mult : gcd (m, k ∗ m + n) = gcd (m, n)apply (induct k)apply (simp-all add : gcd-add2 add-assoc)done

More multiplication laws

lemma gcd-mult-cancel : gcd (k , n) = 1 ==> gcd (k ∗ m, n) = gcd (m, n)apply (rule dvd-anti-sym)apply (rule gcd-greatest)apply (rule-tac n = k in relprime-dvd-mult)apply (simp add : gcd-assoc)apply (simp add : gcd-commute)apply (simp-all add : mult-commute gcd-dvd1 gcd-dvd2 )

apply (blast intro: gcd-dvd1 dvd-trans)done

end

7 Permutations Lawrence C Paulson and ThomasM Rasmussen

theory Permutation = Main:

constsperm :: ( ′a list ∗ ′a list) set

syntax-perm :: ′a list => ′a list => bool (- <∼∼> - [50 , 50 ] 50 )

translationsx <∼∼> y == (x , y) ∈ perm

inductive permintros

Nil [intro!]: [] <∼∼> []swap [intro!]: y # x # l <∼∼> x # y # lCons [intro!]: xs <∼∼> ys ==> z # xs <∼∼> z # ys

38

Page 39: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

trans [intro]: xs <∼∼> ys ==> ys <∼∼> zs ==> xs <∼∼> zs

lemma perm-refl [iff ]: l <∼∼> lapply (induct l)apply autodone

7.1 Some examples of rule induction on permutations

lemma xperm-empty-imp-aux : xs <∼∼> ys ==> xs = [] −−> ys = []— the form of the premise lets the induction bind xs and ys

apply (erule perm.induct)apply (simp-all (no-asm-simp))

done

lemma xperm-empty-imp: [] <∼∼> ys ==> ys = []apply (insert xperm-empty-imp-aux )apply blastdone

This more general theorem is easier to understand!

lemma perm-length: xs <∼∼> ys ==> length xs = length ysapply (erule perm.induct)

apply simp-alldone

lemma perm-empty-imp: [] <∼∼> xs ==> xs = []apply (drule perm-length)apply autodone

lemma perm-sym: xs <∼∼> ys ==> ys <∼∼> xsapply (erule perm.induct)

apply autodone

lemma perm-mem [rule-format ]: xs <∼∼> ys ==> x mem xs −−> x mem ysapply (erule perm.induct)

apply autodone

7.2 Ways of making new permutations

We can insert the head anywhere in the list.

lemma perm-append-Cons: a # xs @ ys <∼∼> xs @ a # ysapply (induct xs)apply autodone

39

Page 40: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma perm-append-swap: xs @ ys <∼∼> ys @ xsapply (induct xs)apply simp-all

apply (blast intro: perm-append-Cons)done

lemma perm-append-single: a # xs <∼∼> xs @ [a]apply (rule perm.trans)prefer 2apply (rule perm-append-swap)apply simpdone

lemma perm-rev : rev xs <∼∼> xsapply (induct xs)apply simp-allapply (blast intro!: perm-append-single intro: perm-sym)done

lemma perm-append1 : xs <∼∼> ys ==> l @ xs <∼∼> l @ ysapply (induct l)apply autodone

lemma perm-append2 : xs <∼∼> ys ==> xs @ l <∼∼> ys @ lapply (blast intro!: perm-append-swap perm-append1 )done

7.3 Further results

lemma perm-empty [iff ]: ([] <∼∼> xs) = (xs = [])apply (blast intro: perm-empty-imp)done

lemma perm-empty2 [iff ]: (xs <∼∼> []) = (xs = [])apply autoapply (erule perm-sym [THEN perm-empty-imp])done

lemma perm-sing-imp [rule-format ]: ys <∼∼> xs ==> xs = [y ] −−> ys = [y ]apply (erule perm.induct)

apply autodone

lemma perm-sing-eq [iff ]: (ys <∼∼> [y ]) = (ys = [y ])apply (blast intro: perm-sing-imp)done

40

Page 41: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma perm-sing-eq2 [iff ]: ([y ] <∼∼> ys) = (ys = [y ])apply (blast dest : perm-sym)done

7.4 Removing elements

constsremove :: ′a => ′a list => ′a list

primrecremove x [] = []remove x (y # ys) = (if x = y then ys else y # remove x ys)

lemma perm-remove: x ∈ set ys ==> ys <∼∼> x # remove x ysapply (induct ys)apply autodone

lemma remove-commute: remove x (remove y l) = remove y (remove x l)apply (induct l)apply autodone

Congruence rule

lemma perm-remove-perm: xs <∼∼> ys ==> remove z xs <∼∼> remove z ysapply (erule perm.induct)

apply autodone

lemma remove-hd [simp]: remove z (z # xs) = xsapply autodone

lemma cons-perm-imp-perm: z # xs <∼∼> z # ys ==> xs <∼∼> ysapply (drule-tac z = z in perm-remove-perm)apply autodone

lemma cons-perm-eq [iff ]: (z#xs <∼∼> z#ys) = (xs <∼∼> ys)apply (blast intro: cons-perm-imp-perm)done

lemma append-perm-imp-perm: !!xs ys. zs @ xs <∼∼> zs @ ys ==> xs <∼∼>ysapply (induct zs rule: rev-induct)apply (simp-all (no-asm-use))apply blastdone

lemma perm-append1-eq [iff ]: (zs @ xs <∼∼> zs @ ys) = (xs <∼∼> ys)

41

Page 42: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (blast intro: append-perm-imp-perm perm-append1 )done

lemma perm-append2-eq [iff ]: (xs @ zs <∼∼> ys @ zs) = (xs <∼∼> ys)apply (safe intro!: perm-append2 )apply (rule append-perm-imp-perm)apply (rule perm-append-swap [THEN perm.trans])

— the previous step helps this blast call succeed quicklyapply (blast intro: perm-append-swap)done

end

8 Fundamental Theorem of Arithmetic (unique fac-torization into primes)

theory Factorization = Primes + Permutation:

8.1 Definitions

constsprimel :: nat list => boolnondec :: nat list => boolprod :: nat list => natoinsert :: nat => nat list => nat listsort :: nat list => nat list

defsprimel-def : primel xs == set xs ⊆ prime

primrecnondec [] = Truenondec (x # xs) = (case xs of [] => True | y # ys => x ≤ y ∧ nondec xs)

primrecprod [] = Suc 0prod (x # xs) = x ∗ prod xs

primrecoinsert x [] = [x ]oinsert x (y # ys) = (if x ≤ y then x # y # ys else y # oinsert x ys)

primrecsort [] = []sort (x # xs) = oinsert x (sort xs)

42

Page 43: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

8.2 Arithmetic

lemma one-less-m: (m::nat) 6= m ∗ k ==> m 6= Suc 0 ==> Suc 0 < mapply (case-tac m)apply autodone

lemma one-less-k : (m::nat) 6= m ∗ k ==> Suc 0 < m ∗ k ==> Suc 0 < kapply (case-tac k)apply autodone

lemma mult-left-cancel : (0 ::nat) < k ==> k ∗ n = k ∗ m ==> n = mapply autodone

lemma mn-eq-m-one: (0 ::nat) < m ==> m ∗ n = m ==> n = Suc 0apply (case-tac n)apply autodone

lemma prod-mn-less-k :(0 ::nat) < n ==> 0 < k ==> Suc 0 < m ==> m ∗ n = k ==> n < k

apply (induct m)apply autodone

8.3 Prime list and product

lemma prod-append : prod (xs @ ys) = prod xs ∗ prod ysapply (induct xs)apply (simp-all add : mult-assoc)done

lemma prod-xy-prod :prod (x # xs) = prod (y # ys) ==> x ∗ prod xs = y ∗ prod ys

apply autodone

lemma primel-append : primel (xs @ ys) = (primel xs ∧ primel ys)apply (unfold primel-def )apply autodone

lemma prime-primel : n ∈ prime ==> primel [n] ∧ prod [n] = napply (unfold primel-def )apply autodone

lemma prime-nd-one: p ∈ prime ==> ¬ p dvd Suc 0

43

Page 44: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (unfold prime-def dvd-def )apply autodone

lemma hd-dvd-prod : prod (x # xs) = prod ys ==> x dvd (prod ys)apply (unfold dvd-def )apply (rule exI )apply (rule sym)apply simpdone

lemma primel-tl : primel (x # xs) ==> primel xsapply (unfold primel-def )apply autodone

lemma primel-hd-tl : (primel (x # xs)) = (x ∈ prime ∧ primel xs)apply (unfold primel-def )apply autodone

lemma primes-eq : p ∈ prime ==> q ∈ prime ==> p dvd q ==> p = qapply (unfold prime-def )apply autodone

lemma primel-one-empty : primel xs ==> prod xs = Suc 0 ==> xs = []apply (unfold primel-def prime-def )apply (case-tac xs)apply simp-alldone

lemma prime-g-one: p ∈ prime ==> Suc 0 < papply (unfold prime-def )apply autodone

lemma prime-g-zero: p ∈ prime ==> 0 < papply (unfold prime-def )apply autodone

lemma primel-nempty-g-one [rule-format ]:primel xs −−> xs 6= [] −−> Suc 0 < prod xs

apply (unfold primel-def prime-def )apply (induct xs)apply (auto elim: one-less-mult)done

44

Page 45: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma primel-prod-gz : primel xs ==> 0 < prod xsapply (unfold primel-def prime-def )apply (induct xs)apply autodone

8.4 Sorting

lemma nondec-oinsert [rule-format ]: nondec xs −−> nondec (oinsert x xs)apply (induct xs)apply (case-tac [2 ] list)apply (simp-all cong del : list .weak-case-cong)

done

lemma nondec-sort : nondec (sort xs)apply (induct xs)apply simp-allapply (erule nondec-oinsert)done

lemma x-less-y-oinsert : x ≤ y ==> l = y # ys ==> x # l = oinsert x lapply simp-alldone

lemma nondec-sort-eq [rule-format ]: nondec xs −−> xs = sort xsapply (induct xs)apply safeapply simp-allapply (case-tac list)apply simp-all

apply (case-tac list)apply simpapply (rule-tac y = aa and ys = lista in x-less-y-oinsert)apply simp-alldone

lemma oinsert-x-y : oinsert x (oinsert y l) = oinsert y (oinsert x l)apply (induct l)apply autodone

8.5 Permutation

lemma perm-primel [rule-format ]: xs <∼∼> ys ==> primel xs −−> primel ysapply (unfold primel-def )apply (erule perm.induct)

apply simp-alldone

lemma perm-prod [rule-format ]: xs <∼∼> ys ==> prod xs = prod ys

45

Page 46: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (erule perm.induct)apply (simp-all add : mult-ac)

done

lemma perm-subst-oinsert : xs <∼∼> ys ==> oinsert a xs <∼∼> oinsert a ysapply (erule perm.induct)

apply autodone

lemma perm-oinsert : x # xs <∼∼> oinsert x xsapply (induct xs)apply autodone

lemma perm-sort : xs <∼∼> sort xsapply (induct xs)apply (auto intro: perm-oinsert elim: perm-subst-oinsert)done

lemma perm-sort-eq : xs <∼∼> ys ==> sort xs = sort ysapply (erule perm.induct)

apply (simp-all add : oinsert-x-y)done

8.6 Existence

lemma ex-nondec-lemma:primel xs ==> ∃ ys. primel ys ∧ nondec ys ∧ prod ys = prod xs

apply (blast intro: nondec-sort perm-prod perm-primel perm-sort perm-sym)done

lemma not-prime-ex-mk :Suc 0 < n ∧ n /∈ prime ==>∃m k . Suc 0 < m ∧ Suc 0 < k ∧ m < n ∧ k < n ∧ n = m ∗ k

apply (unfold prime-def dvd-def )apply (auto intro: n-less-m-mult-n n-less-n-mult-m one-less-m one-less-k)done

lemma split-primel :primel xs ==> primel ys ==> ∃ l . primel l ∧ prod l = prod xs ∗ prod ys

apply (rule exI )apply safeapply (rule-tac [2 ] prod-append)apply (simp add : primel-append)done

lemma factor-exists [rule-format ]: Suc 0 < n −−> (∃ l . primel l ∧ prod l = n)apply (induct n rule: nat-less-induct)apply (rule impI )

46

Page 47: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (case-tac n ∈ prime)apply (rule exI )apply (erule prime-primel)apply (cut-tac n = n in not-prime-ex-mk)apply (auto intro!: split-primel)done

lemma nondec-factor-exists: Suc 0 < n ==> ∃ l . primel l ∧ nondec l ∧ prod l =napply (erule factor-exists [THEN exE ])apply (blast intro!: ex-nondec-lemma)done

8.7 Uniqueness

lemma prime-dvd-mult-list [rule-format ]:p ∈ prime ==> p dvd (prod xs) −−> (∃m. m:set xs ∧ p dvd m)

apply (induct xs)apply (force simp add : prime-def )apply (force dest : prime-dvd-mult)done

lemma hd-xs-dvd-prod :primel (x # xs) ==> primel ys ==> prod (x # xs) = prod ys

==> ∃m. m ∈ set ys ∧ x dvd mapply (rule prime-dvd-mult-list)apply (simp add : primel-hd-tl)apply (erule hd-dvd-prod)done

lemma prime-dvd-eq : primel (x # xs) ==> primel ys ==> m ∈ set ys ==> xdvd m ==> x = mapply (rule primes-eq)apply (auto simp add : primel-def primel-hd-tl)

done

lemma hd-xs-eq-prod :primel (x # xs) ==>

primel ys ==> prod (x # xs) = prod ys ==> x ∈ set ysapply (frule hd-xs-dvd-prod)apply auto

apply (drule prime-dvd-eq)apply auto

done

lemma perm-primel-ex :primel (x # xs) ==>

primel ys ==> prod (x # xs) = prod ys ==> ∃ l . ys <∼∼> (x # l)apply (rule exI )

47

Page 48: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

apply (rule perm-remove)apply (erule hd-xs-eq-prod)apply simp-alldone

lemma primel-prod-less:primel (x # xs) ==>

primel ys ==> prod (x # xs) = prod ys ==> prod xs < prod ysapply (auto intro: prod-mn-less-k prime-g-one primel-prod-gz simp add : primel-hd-tl)done

lemma prod-one-empty :primel xs ==> p ∗ prod xs = p ==> p ∈ prime ==> xs = []

apply (auto intro: primel-one-empty simp add : prime-def )done

lemma uniq-ex-aux :∀m. m < prod ys −−> (∀ xs ys. primel xs ∧ primel ys ∧

prod xs = prod ys ∧ prod xs = m −−> xs <∼∼> ys) ==>primel list ==> primel x ==> prod list = prod x ==> prod x < prod ys==> x <∼∼> list

apply simpdone

lemma factor-unique [rule-format ]:∀ xs ys. primel xs ∧ primel ys ∧ prod xs = prod ys ∧ prod xs = n−−> xs <∼∼> ys

apply (induct n rule: nat-less-induct)apply safeapply (case-tac xs)apply (force intro: primel-one-empty)apply (rule perm-primel-ex [THEN exE ])

apply simp-allapply (rule perm.trans [THEN perm-sym])apply assumptionapply (rule perm.Cons)apply (case-tac x = [])apply (simp add : perm-sing-eq primel-hd-tl)apply (rule-tac p = a in prod-one-empty)apply simp-all

apply (erule uniq-ex-aux )apply (auto intro: primel-tl perm-primel simp add : primel-hd-tl)

apply (rule-tac k = a and n = prod list and m = prod x in mult-left-cancel)apply (rule-tac [3 ] x = a in primel-prod-less)apply (rule-tac [2 ] prod-xy-prod)apply (rule-tac [2 ] s = prod ys in HOL.trans)apply (erule-tac [3 ] perm-prod)apply (erule-tac [5 ] perm-prod [symmetric])apply (auto intro: perm-primel prime-g-zero)

48

Page 49: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

done

lemma perm-nondec-unique:xs <∼∼> ys ==> nondec xs ==> nondec ys ==> xs = ys

apply (rule HOL.trans)apply (rule HOL.trans)apply (erule nondec-sort-eq)apply (erule perm-sort-eq)apply (erule nondec-sort-eq [symmetric])done

lemma unique-prime-factorization [rule-format ]:∀n. Suc 0 < n −−> (∃ !l . primel l ∧ nondec l ∧ prod l = n)

apply safeapply (erule nondec-factor-exists)apply (rule perm-nondec-unique)apply (rule factor-unique)apply simp-all

done

end

9 The set of prime numbers is infinite

theory Prime-number-theorem = Main + Factorization:

9.1 Definition of factorial and basic facts

consts fact :: nat => nat (-! [1000 ] 999 )

primrec0 ! = 1(Suc n)! = n!∗Suc n

lemma zero-fact-bigger-than-zero: (0 ::nat) < 0 !by simp

lemma suc-n-fact-n-fact : (Suc (n::nat))! < (Suc (Suc n)) !proof (induct-tac n)show (Suc 0 )! < (Suc (Suc 0 ))! by simpshow !!n. (Suc n)! < (Suc (Suc n))! ==> (Suc (Suc n))! < (Suc (Suc (Suc n)))!by forceqed

lemma suc-n-fact-bigger-than-zero: (0 ::nat) < (Suc n)!proof (induct-tac n)show 0 < (Suc 0 ) ! by simpshow !!n. 0 < (Suc n)! ==> 0 < (Suc (Suc n))! by simp

49

Page 50: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

qed

lemma zero-l-n-fact [iff ]: (0 ::nat) < n!proof (case-tac n)show n = 0 ==> 0 < n! by simpshow !!nat . n = Suc nat ==> 0 < n! by (simp only : suc-n-fact-bigger-than-zero)qed

9.2 New definition of division for naturals

A new definition of division is introduced instead of the dvd operator; weseek a definition from which the code extration tool can be used

constdefsdvd-k :: nat => nat => natdvd-k a b == (if a dvd b then (SOME k . a ∗ k = b) else 0 )

constsdvd-cons-old :: nat => nat => nat => nat

primrecdvd-cons-old a b 0 = 0dvd-cons-old a b (Suc n) = (if (b < Suc n) then 0 else (if b = a ∗ (Suc n) then(Suc n) else (dvd-cons-old a b n)))

constsdvd-cons-old-simp :: nat => nat => nat => nat

primrecdvd-cons-old-simp a b 0 = 0dvd-cons-old-simp a b (Suc n) = (if b = a ∗ (Suc n) then (Suc n) else (dvd-cons-old-simpa b n))

The dvd cons old function returns the minimum k such that b = a * k; inthe special case where a does not divide b, it returns also 0; in conclusion,whenever it returns something different from zero, a dvd b

lemma dvd-cons-old-g-0 : assumes (0 ::nat) < a and 0 < b shows ∀ k . 0 ≤dvd-cons-old a b kby (simp add : prems)

lemma a-dvd-cons-old-a: assumes 0 < a shows dvd-cons-old (a::nat) a 1 = 1by (simp add : prems)

lemma a-dvd-cons-old-simp-b-k : assumes a-not-0 : 0 < (a::nat) and k-not-0 : 0< (k ::nat)and b-a-k : b = a ∗ k shows dvd-cons-old-simp a b k = kproof −from prems have EX m. k = Suc m by (simp add : not0-implies-Suc)then obtain m where k-suc-m: k = Suc m by auto

50

Page 51: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from k-suc-m and prems have k = (if b = a ∗ k then k else (dvd-cons-old-simpa b m))by (auto simp add : prems)with prems show ?thesis by simpqed

lemma k-dvd-b-k-le-b: assumes a-not-0 : 0 < (a::nat) and k-not-0 : 0 < (k ::nat)and b-a-k : b = a ∗ k shows k ≤ bproof (simp add : prems)from prems show Suc 0 ≤ a by (simp)qed

lemma a-dvd-cons-old-b-k : assumes a-not-0 : 0 < (a::nat) and k-not-0 : 0 <(k ::nat) and b-a-k : b = a ∗ k shows dvd-cons-old a b k = kproof −from prems have k-l-b: k ≤ b by (simp add : k-dvd-b-k-le-b)from prems have EX m. k = Suc m by (simp add : not0-implies-Suc)then obtain m where k-suc-m: k = Suc m by autofrom k-suc-m and k-l-b and prems have k = (if (b < k) then 0 else (if b = a ∗k then k else (dvd-cons-old a b m)))by (auto simp add : prems)with prems and k-l-b show ?thesis by simpqed

lemma a-dvd-cons-old-b-k-not-0 : assumes b-not-0 : 0 < b and a-not-0 : 0 <(a::nat) and a-dvd-b: a dvd b shows ∃ k . 0 < dvd-cons-old a b kproof −from a-dvd-b obtain k where b-a-k : b = a ∗ k using dvd-def by autofrom prems have dvd-cons-old-a-b-k : dvd-cons-old a b k = k by (simp add :

a-dvd-cons-old-b-k)from prems have k-not-0 : 0 6= k by simpfrom k-not-0 and dvd-cons-old-a-b-k have dvd-cons-old-a-b-k-not-0 : 0 < dvd-cons-old

a b k by simpthen show ?thesis by (rule exI )

qed

lemma a-not-dvd-cons-old-b-k : assumes b-not-0 : 0 < b and a-not-0 : 0 < (a::nat)and a-not-dvd-b: ¬ a dvd b shows dvd-cons-old a b k = 0proof (induct-tac k , simp)fix nassume dvd-cons-old a b n = 0show dvd-cons-old a b (Suc n) = 0proof −from a-not-dvd-b have ¬ (∃ k . b = a ∗ k) by (simp add : dvd-def )then have suc-n-not-dvd-b: b 6= a ∗ (Suc n) by fasthave dvd-cons-old a b (Suc n) = (if (b < Suc n) then 0 else (if b = a ∗ (Suc

n) then (Suc n) else (dvd-cons-old a b n))) by simpalso from prems and suc-n-not-dvd-b have . . . = 0by (split split-if )+ (simp)

51

Page 52: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

finally show ?thesis by simpqed

qed

The definition of dvd cons old showed to be not compatible with the codeextraction tool fom the Isabelle code generator

9.3 The infinity of the set of primes using Isabelle definitions

The following lemmas are needed to produce a proof of the Euclid’s Lemmaand the infinity of the set ”primes”

lemma dvd-prod [iff ]: n dvd prod (n # ns)by simp

lemma prime-factor-exists-first-attempt : (1 ::nat) < n ==> ∃ p ∈ prime. p dvdnproof (case-tac n ∈ prime)show [| 1 < n; n : prime |] ==> EX p:prime. p dvd nby (simp add : dvd-def , force)

show [| 1 < n; n ∼: prime |] ==> EX p:prime. p dvd nproof −assume N : 1 < nassume N-not-prime: n /∈ primefrom prems have Existence1 : EX p:prime. p dvd nproof −from N obtain l where primel l ∧ nondec l ∧ prod l = nusing nondec-factor-exists by auto

with N have Existence: EX p:prime. p dvd nproof −from prems have l 6= []by (auto simp add : primel-nempty-g-one)

then have ∃ x xs. l = (x # xs)by (auto simp add : neq-Nil-conv)

with prems show ?thesis by (auto simp add : primel-def )qedfrom prems and Existence show ?thesis by simp

qedfrom prems and Existence1 show ?thesis by simp

qedqed

lemma prime-factor-exists: assumes N : (1 ::nat) < n shows ∃ p ∈ prime. p dvdnproof −from N obtain l where primel l ∧ nondec l ∧ prod l = nusing nondec-factor-exists by auto

with N have Existence: EX p:prime. p dvd nproof −

52

Page 53: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from prems have l 6= []by (auto simp add : primel-nempty-g-one)

then have ∃ x xs. l = (x # xs)by (auto simp add : neq-Nil-conv)

with prems show ?thesis by (auto simp add : primel-def )qedfrom prems and Existence show ?thesis by simp

qed

lemma dvd-factorial-n: assumes N : 0 < n shows n dvd n!proof −from prems have EX m. n = Suc m by (simp add : not0-implies-Suc)then obtain m where n = Suc m by autothen have (n dvd n!) == (Suc m dvd (Suc m) !) by simphave Suc m dvd (Suc m) !proof (induct m, simp)fix nassume Suc-dvd : Suc n dvd (Suc n)!show Suc (Suc n) dvd (Suc (Suc n)) !proof −have (Suc (Suc n)) ! = (Suc (Suc n))∗(Suc n)!by simp

also have Suc (Suc n) dvd (Suc (Suc n)) ∗ (Suc n)!by force

finally show ?thesis by simpqed

qedwith prems show ?thesis by simpqed

lemma m-dvd-n-fact-m-dvd-suc-n-fact : assumes m dvd n! shows m dvd (Suc n)!proof (unfold dvd-def )show ∃ k . (Suc n)! = m ∗ kproof −from prems have EX k . n! = m ∗ k by (auto simp add : dvd-def )then obtain k where N : n! = m ∗ k by autohave (Suc n)! = Suc n ∗ (n!) by simpalso have . . . = Suc n ∗ (k ∗ m) by (simp add : N )also have . . . = Suc n ∗ (m ∗ k)by (simp add : nat-mult-commute)

also have . . . = (Suc n ∗ m) ∗ kby (simp only : nat-mult-assoc)

also have . . . = (m ∗ Suc n) ∗ kby (simp add : nat-mult-commute)

also have . . . = m ∗ (Suc n ∗ k)by (simp only : nat-mult-assoc)

finally show EX k . (Suc n)! = m ∗ k by autoqed

qed

53

Page 54: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma finite-nat-set-is-bounded : finite M ==> EX n::nat . ∀ m ∈ M . m ≤ napply (frule finite-nat-bounded)apply clarifyapply (rule-tac x = k in exI )apply autodone

lemma dvd-factorial [rule-format ]: (0 ::nat) < m −→ m ≤ n −→ m dvd n!proof (induct-tac n, simp, clarify)fix nassume 0 < massume m ≤ Suc nassume m ≤ n −→ m dvd n!show m dvd (Suc n)!proof (cases m = Suc n)from prems show m = Suc n =⇒ m dvd (Suc n)!by (auto simp only : dvd-factorial-n)

from prems show m 6= Suc n =⇒ m dvd (Suc n)!proof −assume m 6= Suc nshow m dvd (Suc n)!proof −from prems have m ≤ n by simpwith prems have m dvd n! by simpthen show ?thesis by (simp only : m-dvd-n-fact-m-dvd-suc-n-fact)

qedqed

qedqed

9.4 Euclid’s lemma and the set of primes is infinite

lemma Euclid : ∃ p ∈ prime. n < pproof −let ?k = n! + 1obtain p where prime: p ∈ prime and dvd : p dvd ?k thm prime-factor-existsusing prime-factor-exists by auto

have n < pproof −have ¬ p ≤ nproofassume p ≤ nwith prime-g-zero have p dvd n! by (rule dvd-factorial)with dvd have p dvd ?k − n! by (rule dvd-diff )

54

Page 55: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

then have p dvd 1 by simpwith prime show False using prime-nd-one by auto

qedthen show ?thesis by simp

qedfrom this and prime show ?thesis ..

qed

corollary prime-set-infinite: ¬ finite primeusing Euclid by (fastsimp dest !: finite-nat-set-is-bounded simp:le-def )

9.5 Recursive definition of division

This is another possible definition of the dvd operator in order to be ableto later extract code

constsdvd-aux :: nat => nat => nat => bool

primrecdvd-aux a b 0 = Falsedvd-aux a b (Suc n) = (if b = a ∗ (Suc n) then True else dvd-aux a b n)

constdefsdvd-cons:: nat => nat => booldvd-cons a b == (dvd-aux a b b)

The definition prime3 is not good enough to extract code from it, due tothe forall

constdefsprime3 :: nat setprime3 == {p. 1 < p ∧ (∀n<p. (dvd-cons n p) = False)}

constdefsprime-number3 :: nat => nat setprime-number3 n == {p. dvd-cons p n}

9.6 Equivalence between the Isabelle defined dvd and therecursive definition

The two following lemmas show the equivalence between dvd and dvd aux(except for the case a = 0, b = 0, which is not considerated)

lemma a-dvd-b-dvd-aux : assumes a-g-0 : 0 < (a::nat) and b-g-0 : 0 < b anda-dvd-p: a dvd b shows ∃ k≤b. dvd-aux a b k = Trueproof −from prems have exists-k : ∃ k≤b. b = a ∗ k by (auto simp add : dvd-def )with prems obtain k where b = a ∗ k and k ≤ b by autofrom prems have k-g-0 : 0 < k by simp

55

Page 56: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from k-g-0 have ∃ k-1 . k = Suc k-1 by (simp add : not0-implies-Suc)then obtain k-1 where k-1 : k = Suc k-1 by autohave dvd-aux a b (Suc k-1 ) = (if b = a ∗ (Suc k-1 ) then True else dvd-aux a b

k-1 ) by simpwith prems have dvd-aux a b k = True by simpwith prems show ?thesis by (auto)

qed

lemma dvd-aux-a-b-k-a-dvd-b: assumes ∃ k≤b. dvd-aux a b k = True shows advd bproof −from prems obtain k where dvd-aux-a-b-k : dvd-aux a b k = True by autofrom this have k-g-0 : 0 < k by (simp) (cases k = 0 , auto)then have ∃ k-1 . k = Suc k-1 by (simp add : not0-implies-Suc)then obtain k-1 where k-1 : k = Suc k-1 by autowith prems have dvd-aux-a-b-suck-1 : dvd-aux a b (Suc k-1 ) = True by simpfrom dvd-aux-a-b-suck-1 show a dvd bproof (induct k-1 )show one: dvd-aux a b (Suc 0 ) = True =⇒ a dvd bproof −assume dvd-aux-a-b-suc0 : dvd-aux a b (Suc 0 ) = Trueshow a dvd bproof −

from dvd-aux-a-b-suc0 have (if b = a ∗ (Suc 0 ) then True else dvd-aux ab 0 ) = True by simp

then have (if b = a ∗ (Suc 0 ) then True else False) = True by (auto simpadd : if-def dvd-aux .simps)

then have b = a ∗ (Suc 0 ) by (cases b = a ∗ (Suc 0 ))(auto)then show a dvd b by simp

qedqedshow two:

∧n. [[dvd-aux a b (Suc n) = True =⇒ a dvd b; dvd-aux a b (Suc

(Suc n)) = True]] =⇒ a dvd bproof −fix nassume dvd-aux a b (Suc n) = True =⇒ a dvd bassume dvd-aux a b (Suc (Suc n)) = Trueshow a dvd bproof −from prems have (if b = a ∗ (Suc (Suc n)) then True else dvd-aux a b (Suc

n)) = True by simpthen show ?thesisproof (cases (b = a ∗ (Suc (Suc n))) = True)

show [[(if b = a ∗ Suc (Suc n) then True else dvd-aux a b (Suc n)) =True; (b = a ∗ Suc (Suc n)) = True]] =⇒ a dvd b

proof −assume (if b = a ∗ Suc (Suc n) then True else dvd-aux a b (Suc n)) =

Trueassume (b = a ∗ Suc (Suc n)) = True

56

Page 57: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

show a dvd bproof −from prems have b = a ∗ (Suc (Suc n)) by simpthen show ?thesis by (simp only : exI dvd-def )

qedqedshow [[(if b = a ∗ Suc (Suc n) then True else dvd-aux a b (Suc n)) =

True; (b = a ∗ Suc (Suc n)) 6= True]] =⇒ a dvd bproof −assume (if b = a ∗ Suc (Suc n) then True else dvd-aux a b (Suc n)) =

Trueassume (b = a ∗ Suc (Suc n)) 6= Trueshow a dvd bproof −from prems have b 6= a ∗ Suc (Suc n) by simpwith prems have dvd-aux a b (Suc n) = True by simpwith prems show ?thesis by simp

qedqed

qedqed

qedqed (simp-all)

qed

corollary dvd-aux-a-b-b-a-dvd-b: assumes dvd-aux a b b = True shows a dvd bby (rule dvd-aux-a-b-k-a-dvd-b, auto simp add : prems)

corollary dvd-cons-impl-dvd : assumes dvd-cons a b = True shows a dvd bproof −from prems have dvd-aux a b b = True by (unfold dvd-cons-def , simp)then show ?thesis by (rule dvd-aux-a-b-b-a-dvd-b)

qed

lemma dvd-aux-impl-dvd-aux-suc: assumes dvd-aux a b k = True shows dvd-auxa b (Suc k) = Trueproof −have (dvd-aux a b (Suc k) = True) = ((if b = a ∗ (Suc k) then True else dvd-aux

a b k) = True)by simp

also have . . .by (cases b = a ∗ (Suc k)) (auto simp add : prems)

finally show ?thesis by simpqed

lemma dvd-aux-impl-dvd-aux-g [rule-format ]: dvd-aux a b k = True −→ k ≤ n−→ dvd-aux a b n = Trueproof (cases k = n)

show k = n =⇒ dvd-aux a b k = True −→ k ≤ n −→ dvd-aux a b n = True

57

Page 58: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

by simpshow k 6= n =⇒ dvd-aux a b k = True −→ k ≤ n −→ dvd-aux a b n = Trueproof −assume k 6= nshow dvd-aux a b k = True −→ k ≤ n −→ dvd-aux a b n = Trueproof −

from prems have (dvd-aux a b k = True −→ k ≤ n −→ dvd-aux a b n =True)

= (dvd-aux a b k = True −→ k < n −→ dvd-aux a b n = True) by autoalso have . . .proof (induct n)show dvd-aux a b k = True −→ k < 0 −→ dvd-aux a b 0 = True by simpshow

∧n. dvd-aux a b k = True −→ k < n −→ dvd-aux a b n =⇒ dvd-aux

a b k = True −→ k < Suc n −→ dvd-aux a b (Suc n) = Trueproof (auto simp only : dvd-aux-impl-dvd-aux-suc)fix nassume dvd-aux a b kassume ¬ k < nassume k < Suc nshow dvd-aux a b (Suc n)proof −from prems have k = n by simpwith prems have dvd-aux a b n by simpthen show ?thesis by (simp only : dvd-aux-impl-dvd-aux-suc)

qedqed

qed (simp-all)finally show ?thesis by simp

qedqed

qed

lemma dvd-aux-impl-dvd-cons: assumes ∃ k≤b. dvd-aux a b k = True showsdvd-cons a b = Trueproof (unfold dvd-cons-def )show dvd-aux a b b = Trueproof −from prems obtain k where dvd-aux-a-b-k : dvd-aux a b k = True and k-le-b:

k ≤ b by autothen show dvd-aux a b b = Trueby (rule dvd-aux-impl-dvd-aux-g)

qedqed

corollary dvd-impl-dvd-cons: assumes a-g-0 : 0 < (a::nat) and b-g-0 : 0 < b anda-dvd-p: a dvd b shows dvd-cons a b = Trueproof −from prems have exists-k : ∃ k . b = a ∗ k by (simp add : dvd-def )then obtain k where b = a ∗ k by auto

58

Page 59: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from prems have suc-0-le-b: Suc 0 ≤ b by simpwith prems and exists-k have k-le-b: k ≤ b by (simp)from prems have k-g-0 : 0 < k by simpfrom k-g-0 have ∃ k-1 . k = Suc k-1 by (simp add : not0-implies-Suc)then obtain k-1 where k-1 : k = Suc k-1 by autohave dvd-aux a b (Suc k-1 ) = (if b = a ∗ (Suc k-1 ) then True else dvd-aux a b

k-1 ) by simpwith prems have dvd-aux-a-b-k : dvd-aux a b k = True by simpfrom dvd-aux-a-b-k and k-le-b have ∃ k ≤ b. dvd-aux a b k = True by autothen show ?thesis by (simp only : dvd-aux-impl-dvd-cons)

qed

lemma dvd-equiv-dvd-cons: assumes 0 < a and 0 < b shows a dvd b = dvd-consa bproof −from prems have one: a dvd b =⇒ dvd-cons a bby (simp only : dvd-impl-dvd-cons)have two: dvd-cons a b =⇒ a dvd b by (simp add : dvd-cons-impl-dvd)from one and two show ?thesis by fastqed

We have proved that dvd implies dvd cons and that dvd cons implies dvd

The following step is to give a constructive definition for the prime numbers

9.7 Definition of prime number recursively

The function prime aux is recursive and returns False if p is not prime, andTrue if there is an n such that n dvd p

constsprime-aux :: nat => nat => bool

primrecprime-aux p 0 = Falseprime-aux p (Suc n) = (if (p = 1 ) then False else (if (n = 0 ∨ n = 1 ) then Trueelse (if (dvd-cons n p = True) then False else prime-aux p n)))

The function prime cons turns prime aux into a function of one argument

constsprime-cons :: nat => bool

defsprime-cons-def : prime-cons x == prime-aux x x

lemma prime-aux-p-0-true [simp]: shows prime-aux p 0 = Falseby auto

59

Page 60: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma prime-aux-p-1-true [simp]: assumes 1 < p shows prime-aux p 1 = True

proof −from prems show ?thesis by auto

qed

lemma prime-aux-1-1 : shows prime-aux 1 1 = False by auto

lemma prime-cons-1 [simp]: shows prime-cons 1 = False by (simp add : prime-cons-def )

9.8 Equivalence between the Isabelle prime definition andour recursive one

lemma prime-aux-suc-impl-prime-aux : assumes 1 < k and prime-aux p (Suc k)= True shows prime-aux p k = Trueproof −from prems have (prime-aux p (Suc k) = True)= ((if (p = 1 ) then False else (if (k = 0 ∨ k = 1 ) then True else (if (dvd-cons

k p = True) then False else prime-aux p k))) = True)by simp

also from prems have . . . = ((if (k = 0 ∨ k = 1 ) then True else (if (dvd-consk p = True) then False else prime-aux p k)) = True)proof (cases p = 1 )case Truefrom True have prime-aux p (Suc k) = False by simpwith prems show ?thesis by simp

nextcase Falsefrom False show ?thesis by simp

qedalso from prems have . . . = ((if dvd-cons k p = True then False else prime-aux

p k) = True)by force

also have . . . = (prime-aux p k = True)proof (cases dvd-cons k p)case Truefrom True and prems have prime-aux p (Suc k) = False by simpwith prems show ?thesis by simp

nextcase Falsefrom False show ?thesis by simp

qedfinally show ?thesis using prems by simp

qed

lemma prime-aux-impl-prime-aux-suc: assumes 1 < k and prime-aux p k = Falseshows prime-aux p (Suc k) = Falseproof −have (prime-aux p (Suc k) = False) =

60

Page 61: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

((if (p = 1 ) then False else (if (k = 0 ∨ k = 1 ) then True else (if (dvd-consk p = True) then False else prime-aux p k))) = False)

by simpalso have . . .proof (cases p = 1 )case Truefrom True show ?thesis by simp

nextcase Falsefrom False show ?thesisproof −

from False have ((if p = 1 then False else if (k = 0 ∨ k = 1 ) then Trueelse if dvd-cons k p = True then False else prime-aux p k) = False) =

((if (k = 0 ∨ k = 1 ) then True else if dvd-cons k p = True then False elseprime-aux p k) = False) by simp

also have . . .proof (cases k = 0 ∨ k = 1 )case Truefrom True and prems show ?thesis by simp

nextcase Falsefrom False have ((if (k = 0 ∨ k = 1 ) then True else if dvd-cons k p =

True then False else prime-aux p k) = False) =((if dvd-cons k p = True then False else prime-aux p k) = False)by simp

also have . . .proof (cases dvd-cons k p = True)case Truefrom True show ?thesis by simp

nextcase Falsefrom False and prems show ?thesis by simp

qedfinally show ?thesis using prems by simp

qedfinally show ?thesis using prems by auto

qedqedfinally show ?thesis by simp

qed

We prove first the following induction rule because it goes upwards; in thenext one, the induction starts from p and goes down

lemma prime-aux-suc-impl-prim-aux : [[ 1 < k ; prime-aux p k = False; k ≤ m]]=⇒ prime-aux p (Suc m) = Falseproof (induct m, simp)fix nassume prime-aux-p-n: [[1 < k ; prime-aux p k = False; k ≤ n]] =⇒ prime-aux

p (Suc n) = False

61

Page 62: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

assume 1 < kassume prime-aux p k = Falseassume k ≤ Suc nshow prime-aux p (Suc (Suc n)) = Falseproof (rule prime-aux-impl-prime-aux-suc)from prems show 1 < Suc n by simpfrom prems show prime-aux p (Suc n) = Falseproof (cases k = Suc n)case Truefrom True and prems show ?thesis by simp

nextcase Falsefrom False and prems have k ≤ n by simpwith prems show ?thesis by simp

qedqed

qed

corollary prime-aux-false-prime-p-false: [[ 1 < k ; prime-aux p k = False; k ≤p]] =⇒ prime-aux p p = Falseproof −assume 1 < kassume prime-aux p k = Falseassume k ≤ pshow prime-aux p p = Falseproof (cases k = p)case Truefrom True and prems show ?thesis by simp

nextcase Falseshow ?thesisproof −from prems have 0 < p by simpthen have ∃ k1 . p = Suc k1 by (simp add : not0-implies-Suc)with prems obtain k1 where p-suc: p = Suc k1 and k-l-suc: k < Suc k1

by autothen have (prime-aux p p = False) = (prime-aux p (Suc k1 ) = False) by

simpalso have . . .proof (rule prime-aux-suc-impl-prim-aux [of k ])from prems show 1 < k by simpfrom prems show prime-aux p k = False by simpfrom prems show k ≤ k1 by simp

qedfinally show ?thesis by simp

qedqed

qed

62

Page 63: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma prime-aux-suc-impl-prim-aux : [[ 1 < k ; k ≤ p; prime-aux p p = True]]=⇒ prime-aux p k = Trueproof (cases prime-aux p k = True)case Truefrom True show ?thesis by simp

nextcase Falseassume k-g-1 : 1 < kassume k-le-p: k ≤ passume prime-aux-p-p: prime-aux p p = Trueshow prime-aux p k = Trueproof −from False and k-g-1 and k-le-p have prime-aux p p = False by (simp add :

prime-aux-false-prime-p-false)with prime-aux-p-p show ?thesis by simp

qedqed

lemma prime-aux-true-not-dvd-cons: [[ 1 < k ; prime-aux p (Suc k) = True ]] =⇒dvd-cons k p = Falseproof (cases dvd-cons k p = False)case Truefrom True show ?thesis by simp

nextcase Falseassume 1 < kassume prime-aux p (Suc k) = Truefrom prems have (prime-aux p (Suc k) = True)

= ((if (k = 0 ∨ k = 1 ) then True else if dvd-cons k p = True then False elseprime-aux p k) = True)

by simpalso from prems have . . . = ((if dvd-cons k p = True then False else prime-aux

p k) = True)by force

also from False have . . . = False by simpfinally have prime-aux p (Suc k) = False by simpwith prems show ?thesis by simp

qed

lemma prime-aux-true-not-dvd-cons-generic: [[ 1 < k ; k < p ; prime-aux p p =True]] =⇒ dvd-cons k p = Falseproof −fix kassume k-g-1 : (1 ::nat) < kassume k-l-p: k < passume prime-aux-p-p: prime-aux p p = Trueshow dvd-cons k p = Falseproof −from k-l-p have suc-k-le-p: (Suc k) ≤ p by simp

63

Page 64: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

have prime-aux-p-suc-k : prime-aux p (Suc k) = Trueproof (rule prime-aux-suc-impl-prim-aux )from prems show 1 < Suc k by forceshow (Suc k) ≤ p by (simp only : suc-k-le-p)from prems show prime-aux p p = True by simp

qedwith k-g-1 show ?thesis by (simp only : prime-aux-true-not-dvd-cons)

qedqed

The condition prime aux p p = True implies that p is prime

lemma prime-aux-is-prime: assumes prime-aux-p-p: prime-aux p p = True showsp ∈ primeproof (unfold prime-def , clarify , intro conjI )show 1 < pproof −have prime-aux-0 : prime-aux 0 0 = False by autowith prems have p-g-0 : 0 6= p by forcethen have p-g-0 : 0 < p by simphave prime-aux-1 : prime-aux 1 1 = False by autowith prems have 1 6= p by forcewith p-g-0 show 1 < p by simp

qedshow ∀m. m dvd p −→ m = 1 | m = pproof (clarify)fix mshow [[ m dvd p; m 6= p]] =⇒ m = 1proof −assume m-dvd-p: m dvd passume m-not-p: m 6= pshow m = 1proof (cases m = 1 )case Falsehave dvd-cons-false: dvd-cons m p = Falseproof (rule prime-aux-true-not-dvd-cons-generic)show 1 < mproof (cases 0 = m ∨ 1 = m)case Truehave prime-aux-0 : prime-aux 0 0 = False by autowith prems have p-g-0 : 0 6= p by forcethen have p-g-0 : 0 < p by simphave prime-aux-1 : prime-aux 1 1 = False by autowith prems have 1 6= p by forcewith p-g-0 have p-not-1 : 1 < p by simpfrom m-dvd-p have ∃ k . p = m ∗ k by (simp add : dvd-def )then obtain k where p-m-k : p = m ∗ k by autofrom True show ?thesisproof (cases m = 0 )case True

64

Page 65: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from True and p-m-k and p-not-1 show ?thesis by simpnextcase Falsefrom False and prems show ?thesis by simp

qednextcase Falsefrom False show ?thesis by force

qedshow m < pproof −from m-dvd-p have m ≤ pproof (rule dvd-imp-le)have prime-aux-0 : prime-aux 0 0 = False by autowith prems have p-g-0 : 0 6= p by forcethen show 0 < p by simp

qedwith prems show ?thesis by simp

qedshow prime-aux p p = True by (simp add : prems)

qedthen have ¬ (m dvd p)proof −have m dvd p = dvd-cons m pproof (rule dvd-equiv-dvd-cons)show 0 < pproof −have prime-aux-0 : prime-aux 0 0 = False by autowith prems have p-g-0 : 0 6= p by forcethen show 0 < p by simp

qedshow 0 < mproof −have prime-aux-0 : prime-aux 0 0 = False by autowith prems have p-g-0 : 0 6= p by forcethen have p-g-0 : 0 < p by simpfrom m-dvd-p have ∃ k . p = m ∗ k by (simp add : dvd-def )then obtain k where p = m ∗ k by autowith p-g-0 show ?thesis by simp

qedqedwith dvd-cons-false show ¬ m dvd p by simp

qedwith prems show ?thesis by simp

nextcase Truefrom True show ?thesis by simp

qedqed

65

Page 66: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

qedqed

corollary prime-cons-implies-prime: assumes prime-cons p = True shows p ∈primeproof (rule prime-aux-is-prime)from prems show prime-aux p p = True by (unfold prime-cons-def , simp)qed

lemma prime-prime-aux-suc: assumes k-g-1 : 1 < k and k < p and p-prime: p∈ prime shows prime-aux p (Suc k) = prime-aux p kproof −have prime-aux p (Suc k) =

(if (p = 1 ) then False else (if (k = 0 ∨ k = 1 ) then True else (if (dvd-cons kp = True) then False else prime-aux p k)))by simp

also have . . . = (if (k = 0 ∨ k = 1 ) then True else (if (dvd-cons k p = True)then False else prime-aux p k))proof −from prems have 1 < p by (simp add : prime-def )then show ?thesis by simp

qedalso have . . . = (if (dvd-cons k p = True) then False else prime-aux p k)proof −from k-g-1 have k 6= 0 ∧ k 6= 1 by simpthen show ?thesis by simp

qedalso have . . . = prime-aux p kproof −have dvd-cons k p = False thm dvd-equiv-dvd-consproof −from p-prime have k dvd p =⇒ k = 1 | k = p by (auto simp add : prime-def )with prems have ¬ k dvd p by autowith prems have ¬ dvd-cons k p by (simp add : dvd-equiv-dvd-cons)then show ?thesis by simp

qedthen show ?thesis by simp

qedfinally show ?thesis by simpqed

lemma prime-aux-p-one: assumes 1 < p shows prime-aux p 1 = Trueproof −

have prime-aux p 1 = (if p = 1 then False else (if ((0 ::nat) = 0 ∨ (0 ::nat)= 1 ) then True else (if dvd-cons 0 p = True then False else prime-aux p 0 ))) bysimp

also from prems have . . . = True by autofinally show ?thesis by simp

qed

66

Page 67: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

corollary prime-aux-prime-one-true: assumes p ∈ prime shows prime-aux p 1= Trueproof (rule prime-aux-p-one)from prems show 1 < p by (auto simp add : prime-def )

qed

lemma prime-aux-prime-two-true: assumes p ∈ prime shows prime-aux p (Suc1 ) = Trueproof −have prime-aux p (Suc 1 ) = (if p = 1 then False else (if ((1 ::nat) = 0 ∨ (1 ::nat)

= 1 ) then True else (if dvd-cons 1 p = True then False else prime-aux p 1 )))by simp

also from prems have . . . =(if ((1 ::nat) = 0 ∨ (1 ::nat) = 1 ) then True else(if dvd-cons 1 p = True then False else prime-aux p 1 ))

by (auto simp add : prime-def )also have . . . = Trueby simp

finally show ?thesis by simpqed

lemma prime-prime-aux-two: assumes k-g-1 : 1 < k and k-l-p: k < p and p ∈prime shows prime-aux p (Suc 1 ) = prime-aux p kproof −from k-g-1 and k-l-p show ?thesisproof (induct k , simp)fix nassume [[1 < n; n < p]] =⇒ prime-aux p (Suc 1 ) = prime-aux p nassume 1 < Suc nassume Suc n < pshow prime-aux p (Suc 1 ) = prime-aux p (Suc n)proof (cases Suc n = Suc 1 )case Truefrom True and prems show ?thesis by simp

nextcase Falsefrom prems have n-g-1 : 1 < n by simpfrom prems have n-l-p: n < p by simpfrom prems and n-g-1 and n-l-p have prime-aux p (Suc 1 ) = prime-aux p

n by simpalso have prime-aux p n = prime-aux p (Suc n)proof (rule sym, rule prime-prime-aux-suc)from n-g-1 show 1 < n by simpfrom n-l-p show n < p by simpfrom prems show p ∈ prime by simp

qedfinally show ?thesis by simp

qedqed

67

Page 68: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

qed

corollary prime-implies-prime-aux : assumes p ∈ prime shows prime-aux p p =Trueproof (cases p = Suc 1 )case Truefrom True have prime-aux p p = prime-aux p (Suc 1 ) by simpalso have . . . = Trueproof (rule prime-aux-prime-two-true)from prems show p ∈ prime by simp

qedfinally show ?thesis by simp

nextcase Falsefrom prems have p-g-suc-1 : Suc 1 < p by (auto simp add : prime-def )then have ∃ k . p = Suc k by (auto simp add : not0-implies-Suc)with p-g-suc-1 obtain k where p-suc-k : p = Suc k and k-g-1 : 1 < k and k-l-p:

k < p by autohave prime-aux p (Suc k) = prime-aux p kproof (rule prime-prime-aux-suc)from k-g-1 show 1 < k by (simp)from k-l-p show k < p by simpfrom prems show p ∈ prime by simp

qedwith p-suc-k have prime-aux-p-k : prime-aux p p = prime-aux p k by simpalso have prime-aux p k = prime-aux p (Suc 1 )proof (rule sym, rule prime-prime-aux-two)from k-g-1 show 1 < k by (simp)from k-l-p show k < p by simpfrom prems show p ∈ prime by simp

qedalso have prime-aux p (Suc 1 ) = Trueproof (rule prime-aux-prime-two-true)from prems show p ∈ prime by simp

qedfinally show prime-aux-p-k-prime-aux-two:prime-aux p p = True by simp

qed

corollary prime-implies-prime-cons: assumes p ∈ prime shows prime-cons p =Trueproof (unfold prime-cons-def )show prime-aux p p = True by (rule prime-implies-prime-aux , simp add : prems)qed

corollary prime-equiv-prime-cons: p ∈ prime = prime-cons pproof −have one: p ∈ prime =⇒ prime-cons p by (simp add : prime-implies-prime-cons)have two: prime-cons p =⇒ p ∈ prime by (simp add : prime-cons-implies-prime)from one and two show ?thesis by fast

68

Page 69: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

qed

9.9 Function yielding a prime number smaller than the ar-gument and dividing it

The function some prime divisor returns a prime number smaller than x andsuch that it divides x

constssome-prime-divisor-aux :: nat => nat => nat

primrecsome-prime-divisor-aux x 0 = 0some-prime-divisor-aux x (Suc n) =

(if (prime-cons (Suc n) ∧ dvd-cons (Suc n) x ) then (Suc n)else some-prime-divisor-aux x n)

The function some prime divisor is a simplification of some prime divisor auxin a function of one argument

constssome-prime-divisor :: nat => nat

defssome-prime-divisor-def : some-prime-divisor x == some-prime-divisor-aux x x

The function some big prime returns a prime number smaller than x! + 1;later we will see that it is greater than x

constssome-big-prime :: nat => nat

defssome-big-prime-def : some-big-prime x == some-prime-divisor (x ! + 1 )

constdefsprime-number :: nat => nat setprime-number n == {p. prime-cons p ∧ (dvd-cons p n)}

lemma dvd-cons-and-prime-cons-some-prime-divisor : assumes dvd-cons (Suc k)x and prime-cons (Suc k)shows some-prime-divisor-aux x (Suc k) = Suc k by (auto simp add : prems)

This lemma will help us to do the inductions when unfolding the definitionof some prime divisor aux

lemma some-prime-divisor-aux-bounded : shows some-prime-divisor-aux p k ≤ kby (induct k , auto)

We prove that the prime divisor of a number divides that number

69

Page 70: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

lemma some-prime-divisor-aux-divides-p [rule-format ]: shows (some-prime-divisor-auxp (Suc k) = Suc k) −→ (dvd-cons (Suc k) p) = Trueproof (induct-tac k , simp)fix nassume some-prime-divisor-aux p (Suc n) = Suc n −→ (dvd-cons (Suc n) p) =

Truethen show some-prime-divisor-aux p (Suc (Suc n)) = Suc (Suc n) −→ ( dvd-cons

(Suc (Suc n)) p) = Trueproof (clarify)assume some-prime-divisor-aux p (Suc n) = Suc n −→ dvd-cons (Suc n) p =

Trueassume some-prime-divisor-suc-suc: some-prime-divisor-aux p (Suc (Suc n))

= Suc (Suc n)show dvd-cons (Suc (Suc n)) p = Trueproof −

have (if (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p) then (Suc(Suc n)) else some-prime-divisor-aux p (Suc n))

= some-prime-divisor-aux p (Suc (Suc n))by auto

also have . . . = Suc (Suc n)by (simp only : prems)

finally have (if (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p) then(Suc (Suc n)) else some-prime-divisor-aux p (Suc n))

= Suc (Suc n)by simp

then have prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) pproof (cases prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p, simp)

show [[(if prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p then Suc(Suc n)

else some-prime-divisor-aux p (Suc n)) = Suc (Suc n);¬ (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p)]]=⇒ prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p

proof −assume one: ¬ (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p)assume two: (if prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p

then Suc (Suc n)else some-prime-divisor-aux p (Suc n)) = Suc (Suc n)

show prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) pproof −

from some-prime-divisor-suc-suc have Suc (Suc n) = some-prime-divisor-auxp (Suc (Suc n)) by simp

from one and two have (if prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc(Suc n)) p then Suc (Suc n)

else some-prime-divisor-aux p (Suc n)) = some-prime-divisor-aux p(Suc n) by auto

with some-prime-divisor-suc-suc have contradiction: some-prime-divisor-auxp (Suc n) = Suc (Suc n) by simp

have contradiction2 : some-prime-divisor-aux p (Suc n) ≤ Suc n by(simp only : some-prime-divisor-aux-bounded)

70

Page 71: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from contradiction and contradiction2 show ?thesis by simpqed

qedqedthen show ?thesis by simp

qedqed

qed

We prove that the prime divisor of a number is prime

lemma some-prime-divisor-aux-prime-cons [rule-format ]: shows (some-prime-divisor-auxp (Suc k) = Suc k) −→ (prime-cons (Suc k)) = Trueproof (induct-tac k , simp add : prime-cons-def )fix nassume some-prime-divisor-aux p (Suc n) = Suc n −→ (prime-cons (Suc n)) =

Truethen show some-prime-divisor-aux p (Suc (Suc n)) = Suc (Suc n) −→ (prime-cons

(Suc (Suc n))) = Trueproof (clarify)

assume some-prime-divisor-aux p (Suc n) = Suc n −→ prime-cons (Suc n)= True

assume some-prime-divisor-suc-suc: some-prime-divisor-aux p (Suc (Suc n))= Suc (Suc n)

show prime-cons (Suc (Suc n)) = Trueproof −

have (if (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p) then (Suc(Suc n)) else some-prime-divisor-aux p (Suc n))

= some-prime-divisor-aux p (Suc (Suc n))by auto

also have . . . = Suc (Suc n)by (simp only : prems)

finally have (if (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p) then(Suc (Suc n)) else some-prime-divisor-aux p (Suc n))

= Suc (Suc n)by simp

then have prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) pproof (cases prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p, simp)

show [[(if prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p then Suc(Suc n)

else some-prime-divisor-aux p (Suc n)) = Suc (Suc n);¬ (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p)]]=⇒ prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p

proof −assume one: ¬ (prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p)assume two: (if prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) p

then Suc (Suc n)else some-prime-divisor-aux p (Suc n)) = Suc (Suc n)

show prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc (Suc n)) pproof −

71

Page 72: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

from some-prime-divisor-suc-suc have Suc (Suc n) = some-prime-divisor-auxp (Suc (Suc n)) by simp

from one and two have (if prime-cons (Suc (Suc n)) ∧ dvd-cons (Suc(Suc n)) p then Suc (Suc n)

else some-prime-divisor-aux p (Suc n)) = some-prime-divisor-aux p(Suc n) by auto

with some-prime-divisor-suc-suc have contradiction: some-prime-divisor-auxp (Suc n) = Suc (Suc n) by simp

have contradiction2 : some-prime-divisor-aux p (Suc n) ≤ Suc n by(simp only : some-prime-divisor-aux-bounded)

from contradiction and contradiction2 show ?thesis by simpqed

qedqedthen show ?thesis by simp

qedqed

qed

9.10 Properties of someprimedivisor

corollary some-prime-divisor-aux-prime-cons-and-divides: assumes (some-prime-divisor-auxp (Suc k) = Suc k)shows (prime-cons (Suc k) ∧ dvd-cons (Suc k) p) = True

proof −from prems show ?thesisby (auto simp only : some-prime-divisor-aux-prime-cons some-prime-divisor-aux-divides-p)qed

lemma some-prime-divisor-exists: shows ∃ k . some-prime-divisor p = k by (autosimp add : some-prime-divisor-def )

lemma some-prime-divisor-g-0 : assumes 1 < n shows ∃ p. prime-cons p ∧ dvd-consp nproof −have prime-factor : ∃ p ∈ prime. p dvd nby (rule prime-factor-exists, simp only : prems)then obtain p where p-in-prime: p ∈ prime and p-dvd-n: p dvd n by autofrom p-in-prime have prime-cons-p: prime-cons p by (simp only : prime-implies-prime-cons)have dvd-cons-p-n: dvd-cons p n = Trueproof (rule dvd-impl-dvd-cons)from p-in-prime show 0 < p by (auto simp add : prime-def )from prems show 0 < n by simpfrom p-dvd-n show p dvd n by simpqedfrom prime-cons-p and dvd-cons-p-n show ?thesis by (auto simp add : exI )qed

lemma some-prime-divisor-aux-bounded : some-prime-divisor-aux p k ≤ k

72

Page 73: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

proof (induct k , simp)fix nassume some-prime-divisor-n-bound : some-prime-divisor-aux p n ≤ nshow some-prime-divisor-aux p (Suc n) ≤ Suc nproof (cases prime-cons (Suc n) ∧ dvd-cons (Suc n) p)case Truefrom True show ?thesis by auto

nextcase Falsefrom False and some-prime-divisor-n-bound show ?thesis by auto

qedqed

lemma some-prime-divisor-aux-monotonicity : assumes a-l-b: a < b shows some-prime-divisor-auxp a ≤ some-prime-divisor-aux p bproof −from a-l-b show ?thesisproof (induct b, simp)fix nassume a-l-n: a < n =⇒ some-prime-divisor-aux p a ≤ some-prime-divisor-aux

p nassume a-l-suc-n: a < Suc nshow some-prime-divisor-aux p a ≤ some-prime-divisor-aux p (Suc n)proof (cases a = n)case Truefrom True have some-prime-divisor-aux p a ≤ some-prime-divisor-aux p n

by simpalso have . . . ≤ some-prime-divisor-aux p (Suc n)proof (cases prime-cons (Suc n) ∧ dvd-cons (Suc n) p)case Falsefrom False show ?thesis by auto

nextcase Truefrom True have some-prime-divisor-suc-n: some-prime-divisor-aux p (Suc

n) = Suc n by autohave some-prime-divisor-aux p n ≤ n by (rule some-prime-divisor-aux-bounded)

with some-prime-divisor-suc-n show ?thesis by simpqedfinally show ?thesis by simp

nextcase Falsefrom prems have a < n by simpthen have some-prime-divisor-aux p a ≤ some-prime-divisor-aux p n by

(simp add : prems)also have . . . ≤ some-prime-divisor-aux p (Suc n)proof (cases prime-cons (Suc n) ∧ dvd-cons (Suc n) p)case Falsefrom False show ?thesis by auto

next

73

Page 74: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

case Truefrom True have some-prime-divisor-suc-n: some-prime-divisor-aux p (Suc

n) = Suc n by autohave some-prime-divisor-aux p n ≤ n by (rule some-prime-divisor-aux-bounded)

with some-prime-divisor-suc-n show ?thesis by simpqedfinally show ?thesis by auto

qedqed

qed

lemma some-prime-divisor-gt-zero:assumes 1 < nshows 0 < some-prime-divisor n

proof (cases some-prime-divisor n = 0 )case Falsefrom False show ?thesis by simp

nextcase Truethen have some-prime-divisor-aux-n: some-prime-divisor-aux n n = 0by (unfold some-prime-divisor-def )

from prems have exists-prime-cons: ∃m. prime-cons m ∧ dvd-cons m nby (simp add : some-prime-divisor-g-0 )

then obtain m where m-prime-cons: prime-cons m ∧ dvd-cons m n by autothen have m-g-1 : 1 < mproof −from m-prime-cons have m ∈ prime by (simp add : prime-equiv-prime-cons)then show 1 < m by (auto simp add : prime-def )

qedthen have ∃ s. m = Suc s by (auto simp add : not0-implies-Suc)then obtain s where m-suc-s: m = Suc s by autothen have some-prime-divisor-aux n m = some-prime-divisor-aux n (Suc s)by simp

also have . . . = (if (prime-cons (Suc s) ∧ dvd-cons (Suc s) n) then (Suc s) elsesome-prime-divisor-aux n s)

by autoalso from m-suc-s and m-prime-cons have . . . = mby simp

finally have some-prime-divisor-aux-n-m-is-m: some-prime-divisor-aux n m =m by simpfrom some-prime-divisor-aux-n and some-prime-divisor-aux-n-m-is-m and m-g-1have one: some-prime-divisor-aux n n < some-prime-divisor-aux n m by simphave two: some-prime-divisor-aux n m ≤ some-prime-divisor-aux n nproof (cases m = n)case Truefrom True show ?thesis by simp

nextcase Falseshow ?thesis

74

Page 75: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

proof (rule some-prime-divisor-aux-monotonicity)have m-dvd-n: m dvd nproof (rule dvd-cons-impl-dvd)from prems show dvd-cons m n = True by simp

qedhave m ≤ nproof (rule dvd-imp-le)from m-dvd-n show m dvd n by simpfrom prems show 0 < n by simp

qedwith prems show m < n by simp

qedqedfrom one and two show ?thesis by simp

qed

lemma some-prime-divisor-aux-suc-properties [rule-format ]:shows some-prime-divisor-aux k p = a −→ prime-cons a ∧ dvd-cons a k−→ prime-cons (some-prime-divisor-aux k (Suc p)) ∧ dvd-cons (some-prime-divisor-aux

k (Suc p)) kproof (cases prime-cons (Suc p) ∧ dvd-cons (Suc p) k)case Truefrom True show ?thesis by auto

nextcase Falsehave some-prime-divisor-aux k (Suc p) = (if prime-cons (Suc p) ∧ dvd-cons (Suc

p) k then (Suc p) else some-prime-divisor-aux k p)by auto

also from False have . . . = some-prime-divisor-aux k pby auto

finally have some-prime-divisor-aux k (Suc p) = some-prime-divisor-aux k p bysimpwith prems show ?thesis by simp

qed

lemma some-prime-divisor-aux-induction-properties:shows [[ some-prime-divisor-aux k p = a ; prime-cons a ∧ dvd-cons a k ; a ≤ b

]] =⇒prime-cons (some-prime-divisor-aux k b) ∧ dvd-cons (some-prime-divisor-aux k

b) kproof (induct b, simp)fix nassume induct-hypo: [[some-prime-divisor-aux k p = a; prime-cons a ∧ dvd-cons

a k ; a ≤ n]]=⇒ prime-cons (some-prime-divisor-aux k n) ∧ dvd-cons (some-prime-divisor-aux

k n) kassume some-prime-divisor-aux-p-k-is-a: some-prime-divisor-aux k p = a

75

Page 76: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

assume prime-cons-a: prime-cons a ∧ dvd-cons a kassume a-le-suc-n: a ≤ Suc nshow prime-cons (some-prime-divisor-aux k (Suc n)) ∧ dvd-cons (some-prime-divisor-aux

k (Suc n)) kproof (cases a = Suc n)case Truehave some-prime-divisor-aux k (Suc n) = (if (prime-cons (Suc n) ∧ dvd-cons

(Suc n) k) then (Suc n) else some-prime-divisor-aux k n)by auto

also have . . . = Suc nproof −from True and prime-cons-a show ?thesis by simp

qedalso have . . . = a by (simp add : prems)finally have some-prime-divisor-aux k (Suc n) = a by simpwith prime-cons-a show ?thesis by (simp)

nextcase Falsefrom False and prems have a-le-n: a ≤ n by simp

from prems and a-le-n have induct-hypo-impl : prime-cons (some-prime-divisor-auxk n) ∧ dvd-cons (some-prime-divisor-aux k n) k

by simpshow ?thesisproof (cases prime-cons (Suc n) ∧ dvd-cons (Suc n) k)case Truehave some-prime-divisor-aux k (Suc n)

= (if (prime-cons (Suc n) ∧ dvd-cons (Suc n) k) then (Suc n) elsesome-prime-divisor-aux k n)

by autoalso from True have . . . = Suc nby simp

finally have some-prime-divisor-aux k (Suc n) = Suc n by simpwith prems show ?thesis by simp

nextcase Falsehave some-prime-divisor-aux k (Suc n)

= (if (prime-cons (Suc n) ∧ dvd-cons (Suc n) k) then (Suc n) elsesome-prime-divisor-aux k n)

by autoalso from False have . . . = some-prime-divisor-aux k nby auto

finally have some-prime-divisor-aux k (Suc n) = some-prime-divisor-aux kn by simp

with induct-hypo-impl show ?thesis by simpqed

qedqed

lemma some-prime-divisor-properties: assumes x-g-1 : 1 < x shows prime-cons

76

Page 77: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

(some-prime-divisor x ) ∧ dvd-cons (some-prime-divisor x ) xproof −from prems have exists-prime-cons: ∃ p. prime-cons p ∧ dvd-cons p x by (simp

add : some-prime-divisor-g-0 )then obtain p where p-prime-cons: prime-cons p ∧ dvd-cons p x by autothen have p-g-1 : 1 < pproof −from p-prime-cons have p ∈ prime by (simp add : prime-equiv-prime-cons)then show 1 < p by (auto simp add : prime-def )

qedthen have ∃ s. p = Suc s by (auto simp add : not0-implies-Suc)then obtain s where p-suc-s: p = Suc s by autothen have some-prime-divisor-aux x p = some-prime-divisor-aux x (Suc s)by simp

also have . . . = (if (prime-cons (Suc s) ∧ dvd-cons (Suc s) x ) then (Suc s) elsesome-prime-divisor-aux x s)

by autoalso from p-suc-s and p-prime-cons have . . . = pby simp

finally have some-prime-divisor-aux-x-p-is-p: some-prime-divisor-aux x p = pby simphave p-le-x : p ≤ xproof (rule dvd-imp-le)from p-prime-cons show p dvd x by (auto simp add : dvd-cons-impl-dvd)from prems show 0 < x by simp

qedhave prime-cons (some-prime-divisor-aux x x ) ∧ dvd-cons (some-prime-divisor-aux

x x ) xproof (rule some-prime-divisor-aux-induction-properties [of - p p])show some-prime-divisor-aux x p = p by (simp only : some-prime-divisor-aux-x-p-is-p)show prime-cons p ∧ dvd-cons p x by (simp add : p-prime-cons)show p ≤ x by (simp add : p-le-x )

qedthen show ?thesis by (simp add : some-prime-divisor-def )

qed

Here we prove that some prime divisor x divides x

corollary some-prime-divisor-dvd-x : assumes x-g-1 : 1 < x shows some-prime-divisorx dvd xproof −have prime-cons (some-prime-divisor x ) ∧ dvd-cons (some-prime-divisor x ) xby (rule some-prime-divisor-properties, intro prems)then have dvd-cons (some-prime-divisor x ) x by simpthen show some-prime-divisor x dvd xby (simp add : dvd-cons-impl-dvd)

qed

77

Page 78: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

9.11 Euclid’s lemma in a constructive way

For a given number x, there is a prime number p such that p dvd x! + 1and x is smaller than p

theorem x < (some-big-prime x )proof (unfold some-big-prime-def )let ?p = some-prime-divisor (x ! + 1 )from some-prime-divisor-gt-zerohave prime-cons: prime-cons ?p and dvd-cons: dvd-cons ?p (x ! + 1 )by (simp-all add : some-prime-divisor-properties)

from prime-cons have prime-p: ?p ∈ primeby (simp add : prime-equiv-prime-cons)

from dvd-cons have dvd : ?p dvd (x ! + 1 )by (simp add : dvd-cons-impl-dvd)

show x < ?pproof −have ¬ ?p ≤ xproofassume ?p ≤ xwith prime-g-zero and prime-p have ?p dvd x !by (simp add : dvd-factorial)

with dvd have ?p dvd (x ! + 1 ) − x ! by (rule dvd-diff )then have ?p dvd 1 by simpwith prime-p show False using prime-nd-one by auto

qedthen show ?thesis by simp

qedqed

9.12 Code extraction

types-codeset ((- −> bool))

consts-codeCollect ((-))

op : ((- |> -))

generate-code (/home/jesules/Desktop/Prime-number-theorem.ML)fact = factdvd-aux = dvd-auxdvd-cons = dvd-conssome-prime-divisor-aux = some-prime-divisor-auxsome-prime-divisor = some-prime-divisorsome-big-prime = some-big-primeprime-number = prime-number

end

78

Page 79: ProgramExtract - unirioja.es · ProgramExtract Jesus Maria Aransay October 12, 2005 Contents 1 Pi and Function Sets Florian Kammueller and Lawrence C Paulson 1 1.1 Basic Properties

10 Definition of monoid of integers and the rev ofa monoid

theory Monoid = Group:

constdefsrev :: ( ′a, ′m) monoid-scheme => ′a monoidrev G == (| carrier = carrier G , mult = (λx y . mult G y x ), one = one G |)

intg :: int monoidintg == (| carrier = UNIV , mult = (λx y . x + y), one = 0 |)

lemma (in monoid)shows monoid (rev G)apply (unfold rev-def )apply (rule monoidI )apply (auto simp: m-assoc)done

lemma monoid intgapply (unfold intg-def )apply (rule monoidI )apply autodone

10.1 Code extraction

types-codeset ((- −> bool))

consts-codeCollect ((-))op : ((- |> -))

generate-code (/home/jesules/Desktop/Monoid .ML)rev = revintg = intg

end

79