# HG changeset patch # User wenzelm # Date 1139776460 -3600 # Node ID 87cd1ecae3a44c3e47e9a1626788623e3bb1bff3 # Parent 596fb1eb7856ded516f2f04d40400c103a899385 minor tuning of proofs, notably induct; diff -r 596fb1eb7856 -r 87cd1ecae3a4 src/HOL/ex/ThreeDivides.thy --- a/src/HOL/ex/ThreeDivides.thy Sun Feb 12 21:34:18 2006 +0100 +++ b/src/HOL/ex/ThreeDivides.thy Sun Feb 12 21:34:20 2006 +0100 @@ -15,8 +15,8 @@ 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 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 @@ -61,12 +61,12 @@ lemma digit_diff_split: fixes n::nat and nd::nat and x::nat - shows "\n. n = (\x\{.. + shows "n = (\x\{.. (n - (\xx"} *} +text {* Expanding on the previous lemma and lemma @{text "div_sum"}. *} lemma three_divs_1: fixes D :: "nat \ nat" shows "3 dvd (\xnat" shows "3 dvd ((\xxxxxxx (\xxxm. nd = nlen m \ m = (\xx "nlen m" fixing: m) case 0 thus ?case by (simp add: nlen_zero) next case (Suc nd) @@ -184,15 +185,14 @@ 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 obtain c where mexp: "m = 10*(m div 10) + c \ c < 10" .. then have cdef: "c = m mod 10" by arith - show "m = (\xxxxx = (\xxxxxx