| author | wenzelm | 
| Sat, 17 Jun 2006 19:37:53 +0200 | |
| changeset 19912 | 4a3e35fd6e02 | 
| parent 19782 | 48c4632e2c28 | 
| child 19922 | 984ae977f7aa | 
| 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 | 
| 19568 | 9 | imports Main | 
| 10 | begin | |
| 11 | ||
| 12 | section {* Very basic *}
 | |
| 13 | ||
| 14 | consts fib :: "nat \<Rightarrow> nat" | |
| 19736 | 15 | function | 
| 19568 | 16 | "fib 0 = 1" | 
| 17 | "fib (Suc 0) = 1" | |
| 18 | "fib (Suc (Suc n)) = fib n + fib (Suc n)" | |
| 19736 | 19 | by pat_completeness  -- {* shows that the patterns are complete *}
 | 
| 20 |   auto  -- {* shows that the patterns are compatible *}
 | |
| 19568 | 21 | |
| 19736 | 22 | text {* we get partial simp and induction rules: *}
 | 
| 19568 | 23 | thm fib.psimps | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 24 | thm pinduct | 
| 19568 | 25 | |
| 19736 | 26 | text {* There is also a cases rule to distinguish cases along the definition *}
 | 
| 19568 | 27 | thm fib.cases | 
| 28 | ||
| 19736 | 29 | text {* Now termination: *}
 | 
| 19568 | 30 | termination fib | 
| 19736 | 31 | by (auto_term "less_than") | 
| 19568 | 32 | |
| 33 | thm fib.simps | |
| 34 | thm fib.induct | |
| 35 | ||
| 36 | ||
| 37 | section {* Currying *}
 | |
| 38 | ||
| 19736 | 39 | consts add :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 19568 | 40 | function | 
| 41 | "add 0 y = y" | |
| 42 | "add (Suc x) y = Suc (add x y)" | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 43 | thm add_rel.cases | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 44 | |
| 19568 | 45 | by pat_completeness auto | 
| 19736 | 46 | termination | 
| 19568 | 47 | by (auto_term "measure fst") | 
| 48 | ||
| 19736 | 49 | thm add.induct -- {* Note the curried induction predicate *}
 | 
| 19568 | 50 | |
| 51 | ||
| 52 | section {* Nested recursion *}
 | |
| 53 | ||
| 54 | consts nz :: "nat \<Rightarrow> nat" | |
| 55 | function | |
| 56 | "nz 0 = 0" | |
| 57 | "nz (Suc x) = nz (nz x)" | |
| 58 | by pat_completeness auto | |
| 59 | ||
| 19736 | 60 | lemma nz_is_zero: -- {* A lemma we need to prove termination *}
 | 
| 19583 
c5fa77b03442
function-package: Changed record usage to make sml/nj happy...
 krauss parents: 
19568diff
changeset | 61 | assumes trm: "x \<in> nz_dom" | 
| 19568 | 62 | shows "nz x = 0" | 
| 63 | using trm | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 64 | apply induct | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 65 | apply auto | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 66 | done | 
| 19568 | 67 | |
| 68 | termination nz | |
| 19736 | 69 |   apply (auto_term "less_than") -- {* Oops, it left us something to prove *}
 | 
| 19568 | 70 | by (auto simp:nz_is_zero) | 
| 71 | ||
| 72 | thm nz.simps | |
| 73 | thm nz.induct | |
| 74 | ||
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 75 | text {* Here comes McCarthy's 91-function *}
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 76 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 77 | consts f91 :: "nat => nat" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 78 | function | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 79 | "f91 n = (if 100 < n then n - 10 else f91 (f91 (n + 11)))" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 80 | by pat_completeness auto | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 81 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 82 | (* Prove a lemma before attempting a termination proof *) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 83 | lemma f91_estimate: | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 84 | assumes trm: "n : f91_dom" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 85 | shows "n < f91 n + 11" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 86 | using trm by induct auto | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 87 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 88 | (* Now go for termination *) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 89 | termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 90 | proof | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 91 | let ?R = "measure (%x. 101 - x)" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 92 | show "wf ?R" .. | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 93 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 94 | fix n::nat assume "~ 100 < n" (* Inner call *) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 95 | thus "(n + 11, n) : ?R" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 96 | by simp arith | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 97 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 98 | assume inner_trm: "n + 11 : f91_dom" (* Outer call *) | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 99 | with f91_estimate have "n + 11 < f91 (n + 11) + 11" . | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 100 | with `~ 100 < n` show "(f91 (n + 11), n) : ?R" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 101 | by simp arith | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 102 | qed | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 103 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 104 | |
| 19568 | 105 | |
| 106 | section {* More general patterns *}
 | |
| 107 | ||
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 108 | subsection {* Overlapping patterns *}
 | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 109 | |
| 19736 | 110 | text {* Currently, patterns must always be compatible with each other, since
 | 
| 19568 | 111 | no automatich splitting takes place. But the following definition of | 
| 19736 | 112 | gcd is ok, although patterns overlap: *} | 
| 19568 | 113 | |
| 114 | consts gcd2 :: "nat \<Rightarrow> nat \<Rightarrow> nat" | |
| 19736 | 115 | function | 
| 19568 | 116 | "gcd2 x 0 = x" | 
| 117 | "gcd2 0 y = y" | |
| 118 | "gcd2 (Suc x) (Suc y) = (if x < y then gcd2 (Suc x) (y - x) | |
| 119 | else gcd2 (x - y) (Suc y))" | |
| 120 | by pat_completeness auto | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 121 | termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 122 | by (auto_term "measure (\<lambda>(x,y). x + y)") | 
| 19568 | 123 | |
| 124 | thm gcd2.simps | |
| 125 | thm gcd2.induct | |
| 126 | ||
| 127 | ||
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 128 | subsection {* Guards *}
 | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 129 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 130 | 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 | 131 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 132 | consts gcd3 :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 133 | function | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 134 | "gcd3 x 0 = x" | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 135 | "gcd3 0 y = y" | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 136 | "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 | 137 | "\<not> x < y \<Longrightarrow> gcd3 (Suc x) (Suc y) = gcd3 (x - y) (Suc y)" | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 138 | apply (cases xa, case_tac a, auto) | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 139 | apply (case_tac b, auto) | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 140 | done | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 141 | termination | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 142 | by (auto_term "measure (\<lambda>(x,y). x + y)") | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 143 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 144 | thm gcd3.simps | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 145 | thm gcd3.induct | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 146 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 147 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 148 | |
| 19736 | 149 | text {* General patterns allow even strange definitions: *}
 | 
| 19568 | 150 | consts ev :: "nat \<Rightarrow> bool" | 
| 19736 | 151 | function | 
| 19568 | 152 | "ev (2 * n) = True" | 
| 153 | "ev (2 * n + 1) = False" | |
| 19736 | 154 | proof -  -- {* completeness is more difficult here \dots *}
 | 
| 19568 | 155 | assume c1: "\<And>n. x = 2 * n \<Longrightarrow> P" | 
| 156 | and c2: "\<And>n. x = 2 * n + 1 \<Longrightarrow> P" | |
| 157 | have divmod: "x = 2 * (x div 2) + (x mod 2)" by auto | |
| 158 | show "P" | |
| 19736 | 159 | proof cases | 
| 19568 | 160 | assume "x mod 2 = 0" | 
| 161 | with divmod have "x = 2 * (x div 2)" by simp | |
| 162 | with c1 show "P" . | |
| 163 | next | |
| 164 | assume "x mod 2 \<noteq> 0" | |
| 165 | hence "x mod 2 = 1" by simp | |
| 166 | with divmod have "x = 2 * (x div 2) + 1" by simp | |
| 167 | with c2 show "P" . | |
| 168 | qed | |
| 19736 | 169 | qed presburger+   -- {* solve compatibility with presburger *}
 | 
| 19568 | 170 | termination by (auto_term "{}")
 | 
| 171 | ||
| 172 | thm ev.simps | |
| 173 | thm ev.induct | |
| 174 | thm ev.cases | |
| 175 | ||
| 19770 
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 | section {* Mutual Recursion *}
 | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 178 | |
| 
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 | consts | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 181 | evn :: "nat \<Rightarrow> bool" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 182 | od :: "nat \<Rightarrow> bool" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 183 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 184 | function | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 185 | "evn 0 = True" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 186 | "evn (Suc n) = od n" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 187 | and | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 188 | "od 0 = False" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 189 | "od (Suc n) = evn n" | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 190 | by pat_completeness auto | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 191 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 192 | thm evn.psimps | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 193 | thm od.psimps | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 194 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 195 | thm evn_od.pinduct | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 196 | thm evn_od.termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 197 | thm evn_od.domintros | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 198 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 199 | termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 200 | by (auto_term "measure (sum_case (%n. n) (%n. n))") | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 201 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 202 | thm evn.simps | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 203 | thm od.simps | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 204 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 205 | |
| 19736 | 206 | end |