| author | wenzelm | 
| Fri, 19 Jan 2007 22:08:07 +0100 | |
| changeset 22100 | 33d7468302bb | 
| parent 21319 | cf814e36f788 | 
| child 22166 | 0a50d4db234a | 
| permissions | -rw-r--r-- | 
| 19568 | 1 | (* Title: HOL/ex/Fundefs.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Alexander Krauss, TU Muenchen | |
| 4 | ||
| 5 | Examples of function definitions using the new "function" package. | |
| 6 | *) | |
| 7 | ||
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 8 | theory Fundefs | 
| 20528 | 9 | imports Main | 
| 19568 | 10 | begin | 
| 11 | ||
| 12 | section {* Very basic *}
 | |
| 13 | ||
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 14 | fun fib :: "nat \<Rightarrow> nat" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 15 | where | 
| 19568 | 16 | "fib 0 = 1" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 17 | | "fib (Suc 0) = 1" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 18 | | "fib (Suc (Suc n)) = fib n + fib (Suc n)" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 19 | |
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21240diff
changeset | 20 | text {* partial simp and induction rules: *}
 | 
| 19568 | 21 | thm fib.psimps | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 22 | thm fib.pinduct | 
| 19568 | 23 | |
| 19736 | 24 | text {* There is also a cases rule to distinguish cases along the definition *}
 | 
| 19568 | 25 | thm fib.cases | 
| 26 | ||
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 27 | thm fib.domintros | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 28 | |
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21240diff
changeset | 29 | text {* total simp and induction rules: *}
 | 
| 19568 | 30 | thm fib.simps | 
| 31 | thm fib.induct | |
| 32 | ||
| 33 | section {* Currying *}
 | |
| 34 | ||
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 35 | fun add :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 36 | where | 
| 19568 | 37 | "add 0 y = y" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 38 | | "add (Suc x) y = Suc (add x y)" | 
| 19568 | 39 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 40 | thm add.simps | 
| 19736 | 41 | thm add.induct -- {* Note the curried induction predicate *}
 | 
| 19568 | 42 | |
| 43 | ||
| 44 | section {* Nested recursion *}
 | |
| 45 | ||
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 46 | function nz :: "nat \<Rightarrow> nat" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 47 | where | 
| 19568 | 48 | "nz 0 = 0" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 49 | | "nz (Suc x) = nz (nz x)" | 
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 50 | by pat_completeness auto | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 51 | |
| 19736 | 52 | lemma nz_is_zero: -- {* A lemma we need to prove termination *}
 | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20536diff
changeset | 53 | assumes trm: "nz_dom x" | 
| 19568 | 54 | shows "nz x = 0" | 
| 55 | using trm | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 56 | by induct auto | 
| 19568 | 57 | |
| 58 | termination nz | |
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21240diff
changeset | 59 | by (relation "less_than") (auto simp:nz_is_zero) | 
| 19568 | 60 | |
| 61 | thm nz.simps | |
| 62 | thm nz.induct | |
| 63 | ||
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 64 | text {* Here comes McCarthy's 91-function *}
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 65 | |
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20536diff
changeset | 66 | |
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 67 | function f91 :: "nat => nat" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 68 | where | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 69 | "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" | 
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 70 | by pat_completeness auto | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 71 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 72 | (* Prove a lemma before attempting a termination proof *) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 73 | lemma f91_estimate: | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20536diff
changeset | 74 | assumes trm: "f91_dom n" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 75 | shows "n < f91 n + 11" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 76 | using trm by induct auto | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 77 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 78 | termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 79 | proof | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 80 | let ?R = "measure (%x. 101 - x)" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 81 | show "wf ?R" .. | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 82 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 83 | fix n::nat assume "~ 100 < n" (* Inner call *) | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19922diff
changeset | 84 | thus "(n + 11, n) : ?R" by simp | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 85 | |
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20536diff
changeset | 86 | assume inner_trm: "f91_dom (n + 11)" (* Outer call *) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 87 | with f91_estimate have "n + 11 < f91 (n + 11) + 11" . | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19922diff
changeset | 88 | with `~ 100 < n` show "(f91 (n + 11), n) : ?R" by simp | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 89 | qed | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 90 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 91 | |
| 19568 | 92 | |
| 93 | section {* More general patterns *}
 | |
| 94 | ||
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 95 | subsection {* Overlapping patterns *}
 | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 96 | |
| 19736 | 97 | text {* Currently, patterns must always be compatible with each other, since
 | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19922diff
changeset | 98 | no automatic splitting takes place. But the following definition of | 
| 19736 | 99 | gcd is ok, although patterns overlap: *} | 
| 19568 | 100 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 101 | fun gcd2 :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 102 | where | 
| 19568 | 103 | "gcd2 x 0 = x" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 104 | | "gcd2 0 y = y" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 105 | | "gcd2 (Suc x) (Suc y) = (if x < y then gcd2 (Suc x) (y - x) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 106 | else gcd2 (x - y) (Suc y))" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 107 | |
| 19568 | 108 | thm gcd2.simps | 
| 109 | thm gcd2.induct | |
| 110 | ||
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 111 | subsection {* Guards *}
 | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 112 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 113 | text {* We can reformulate the above example using guarded patterns *}
 | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 114 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 115 | function gcd3 :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 116 | where | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 117 | "gcd3 x 0 = x" | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 118 | "gcd3 0 y = y" | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 119 | "x < y \<Longrightarrow> gcd3 (Suc x) (Suc y) = gcd3 (Suc x) (y - x)" | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 120 | "\<not> x < y \<Longrightarrow> gcd3 (Suc x) (Suc y) = gcd3 (x - y) (Suc y)" | 
| 19922 | 121 | apply (case_tac x, case_tac a, auto) | 
| 122 | apply (case_tac ba, auto) | |
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 123 | done | 
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 124 | termination by lexicographic_order | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 125 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 126 | thm gcd3.simps | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 127 | thm gcd3.induct | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 128 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 129 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 130 | text {* General patterns allow even strange definitions: *}
 | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 131 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 132 | function ev :: "nat \<Rightarrow> bool" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 133 | where | 
| 19568 | 134 | "ev (2 * n) = True" | 
| 135 | "ev (2 * n + 1) = False" | |
| 19736 | 136 | proof -  -- {* completeness is more difficult here \dots *}
 | 
| 19922 | 137 | fix P :: bool | 
| 138 | and x :: nat | |
| 19568 | 139 | assume c1: "\<And>n. x = 2 * n \<Longrightarrow> P" | 
| 140 | and c2: "\<And>n. x = 2 * n + 1 \<Longrightarrow> P" | |
| 141 | have divmod: "x = 2 * (x div 2) + (x mod 2)" by auto | |
| 142 | show "P" | |
| 19736 | 143 | proof cases | 
| 19568 | 144 | assume "x mod 2 = 0" | 
| 145 | with divmod have "x = 2 * (x div 2)" by simp | |
| 146 | with c1 show "P" . | |
| 147 | next | |
| 148 | assume "x mod 2 \<noteq> 0" | |
| 149 | hence "x mod 2 = 1" by simp | |
| 150 | with divmod have "x = 2 * (x div 2) + 1" by simp | |
| 151 | with c2 show "P" . | |
| 152 | qed | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 153 | qed presburger+ -- {* solve compatibility with presburger *}
 | 
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 154 | termination by lexicographic_order | 
| 19568 | 155 | |
| 156 | thm ev.simps | |
| 157 | thm ev.induct | |
| 158 | thm ev.cases | |
| 159 | ||
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 160 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 161 | section {* Mutual Recursion *}
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 162 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 163 | fun evn od :: "nat \<Rightarrow> bool" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 164 | where | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 165 | "evn 0 = True" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 166 | | "od 0 = False" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 167 | | "evn (Suc n) = od n" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 168 | | "od (Suc n) = evn n" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 169 | |
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 170 | thm evn.simps | 
| 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 171 | thm od.simps | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 172 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 173 | thm evn_od.pinduct | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 174 | thm evn_od.termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 175 | thm evn_od.domintros | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 176 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 177 | |
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 178 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 179 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 180 | |
| 19736 | 181 | end |