# HG changeset patch # User kleing # Date 1139715078 -3600 # Node ID 0111ecc5490e8ea2864eb78129ee28e032a98ba0 # Parent e1867198df485aaee0f1d9fbea335bd04f6d5ad0 divisibility by 3 theorem, contributed by Benjamin Porter, #85 in http://www.cs.ru.nl/~freek/100/ diff -r e1867198df48 -r 0111ecc5490e src/HOL/Isar_examples/ROOT.ML --- a/src/HOL/Isar_examples/ROOT.ML Sat Feb 11 17:17:55 2006 +0100 +++ b/src/HOL/Isar_examples/ROOT.ML Sun Feb 12 04:31:18 2006 +0100 @@ -14,6 +14,7 @@ time_use_thy "Summation"; time_use_thy "KnasterTarski"; time_use_thy "MutilatedCheckerboard"; +time_use_thy "ThreeDivides"; with_path "../NumberTheory" (no_document time_use_thy) "Primes"; with_path "../NumberTheory" time_use_thy "Fibonacci"; time_use_thy "Puzzle"; diff -r e1867198df48 -r 0111ecc5490e src/HOL/Isar_examples/ThreeDivides.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Isar_examples/ThreeDivides.thy Sun Feb 12 04:31:18 2006 +0100 @@ -0,0 +1,291 @@ +(* Title: HOL/Isar_examples/ThreeDivides.thy + ID: $Id$ + Author: Benjamin Porter, 2005 +*) + +header {* Three Divides Theorem *} + +theory ThreeDivides +imports Main LaTeXsugar +begin + +section {* Abstract *} + +text {* +The following document presents a proof of the Three Divides N theorem +formalised in the Isabelle/Isar theorem proving system. + +{\em Theorem}: 3 divides n if and only if 3 divides the sum of all +digits in n. + +{\em Informal Proof}: +Take $n = \sum{n_j * 10^j}$ where $n_j$ is the $j$'th least +significant digit of the decimal denotation of the number n and the +sum ranges over all digits. Then $$ (n - \sum{n_j}) = \sum{n_j * (10^j +- 1)} $$ We know $\forall j\; 3|(10^j - 1) $ and hence $3|LHS$, +therefore $$\forall n\; 3|n \Longleftrightarrow 3|\sum{n_j}$$ +@{text "\"} +*} + +section {* Formal proof *} + +subsection {* Miscellaneous summation lemmas *} + +text {* We can decompose a sum into an expanded form by removing +its first element. *} + +lemma sum_rmv_head: + fixes m::nat + assumes m: "0 < m" + shows "(\xx\{1.. {0<..nat" + shows + "\x. Q x \ P x \ + (\xxxxxx. a dvd A x \ a dvd (\xxn. n = (\x\{.. + (n - (\xx"} *} +lemma three_divs_1: + fixes D :: "nat \ nat" + shows "3 dvd (\xnat" + shows "3 dvd ((\xxxnat)"}), +we show that 3 divides the expansive sum $\sum{(D\;x)*10^x}$ over $x$ +if and only if 3 divides the sum of the individual numbers +$\sum{D\;x}$. +*} +lemma three_div_general: + fixes D :: "nat \ nat" + shows "(3 dvd (\xxx (\xxxxxxx nat" +recdef nlen "measure id" + "nlen 0 = 0" + "nlen x = 1 + nlen (x div 10)" + +text {* The function @{text "sumdig"} returns the sum of all digits in +some number n. *} + +constdefs + sumdig :: "nat \ nat" + "sumdig n \ \x < nlen n. n div 10^x mod 10" + +text {* Some properties of these functions follow. *} + +lemma nlen_zero: + "0 = nlen x \ x = 0" + by (induct x rule: nlen.induct) auto + +lemma nlen_suc: + "Suc m = nlen n \ m = nlen (n div 10)" + by (induct n rule: nlen.induct) simp_all + + +text {* The following lemma is the principle lemma required to prove +our theorem. It states that an expansion of some natural number $n$ +into a sequence of its individual digits is always possible. *} + +lemma exp_exists: + "\m. nd = nlen m \ m = (\x + m div 10 = (\xc. m = 10*(m div 10) + c \ c < 10" by presburger + from this obtain c where mexp: "m = 10*(m div 10) + c \ c < 10" .. + then have cdef: "c = m mod 10" by arith + show "m = (\xxx = (\x = (\x = (\x\{Suc 0.. = (\xxxxxx