| author | wenzelm | 
| Tue, 18 May 2021 21:09:51 +0200 | |
| changeset 73734 | f7f0d516df0c | 
| parent 70468 | 8406a2c296e0 | 
| permissions | -rw-r--r-- | 
| 62999 | 1 | (* Title: HOL/ex/Functions.thy | 
| 19568 | 2 | Author: Alexander Krauss, TU Muenchen | 
| 22726 | 3 | *) | 
| 19568 | 4 | |
| 61343 | 5 | section \<open>Examples of function definitions\<close> | 
| 19568 | 6 | |
| 62999 | 7 | theory Functions | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
63067diff
changeset | 8 | imports Main "HOL-Library.Monad_Syntax" | 
| 19568 | 9 | begin | 
| 10 | ||
| 61343 | 11 | subsection \<open>Very basic\<close> | 
| 19568 | 12 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 13 | fun fib :: "nat \<Rightarrow> nat" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 14 | where | 
| 19568 | 15 | "fib 0 = 1" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 16 | | "fib (Suc 0) = 1" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 17 | | "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 | 18 | |
| 62999 | 19 | text \<open>Partial simp and induction rules:\<close> | 
| 19568 | 20 | thm fib.psimps | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 21 | thm fib.pinduct | 
| 19568 | 22 | |
| 62999 | 23 | text \<open>There is also a cases rule to distinguish cases along the definition:\<close> | 
| 19568 | 24 | thm fib.cases | 
| 25 | ||
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 26 | |
| 62999 | 27 | text \<open>Total simp and induction rules:\<close> | 
| 19568 | 28 | thm fib.simps | 
| 29 | thm fib.induct | |
| 30 | ||
| 62999 | 31 | text \<open>Elimination rules:\<close> | 
| 53611 | 32 | thm fib.elims | 
| 33 | ||
| 62999 | 34 | |
| 61343 | 35 | subsection \<open>Currying\<close> | 
| 19568 | 36 | |
| 25170 | 37 | fun add | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 38 | where | 
| 19568 | 39 | "add 0 y = y" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 40 | | "add (Suc x) y = Suc (add x y)" | 
| 19568 | 41 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 42 | thm add.simps | 
| 62999 | 43 | thm add.induct \<comment> \<open>Note the curried induction predicate\<close> | 
| 19568 | 44 | |
| 45 | ||
| 61343 | 46 | subsection \<open>Nested recursion\<close> | 
| 19568 | 47 | |
| 62999 | 48 | function nz | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 49 | where | 
| 19568 | 50 | "nz 0 = 0" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 51 | | "nz (Suc x) = nz (nz x)" | 
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 52 | by pat_completeness auto | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 53 | |
| 62999 | 54 | lemma nz_is_zero: \<comment> \<open>A lemma we need to prove termination\<close> | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20536diff
changeset | 55 | assumes trm: "nz_dom x" | 
| 19568 | 56 | shows "nz x = 0" | 
| 57 | using trm | |
| 39754 | 58 | by induct (auto simp: nz.psimps) | 
| 19568 | 59 | |
| 60 | termination nz | |
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21240diff
changeset | 61 | by (relation "less_than") (auto simp:nz_is_zero) | 
| 19568 | 62 | |
| 63 | thm nz.simps | |
| 64 | thm nz.induct | |
| 65 | ||
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 66 | |
| 62999 | 67 | subsubsection \<open>Here comes McCarthy's 91-function\<close> | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20536diff
changeset | 68 | |
| 62999 | 69 | function f91 :: "nat \<Rightarrow> nat" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 70 | where | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 71 | "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 | 72 | by pat_completeness auto | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 73 | |
| 62999 | 74 | text \<open>Prove a lemma before attempting a termination proof:\<close> | 
| 75 | lemma f91_estimate: | |
| 24585 | 76 | assumes trm: "f91_dom n" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 77 | shows "n < f91 n + 11" | 
| 39754 | 78 | using trm by induct (auto simp: f91.psimps) | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 79 | |
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 80 | termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 81 | proof | 
| 62999 | 82 | let ?R = "measure (\<lambda>x. 101 - x)" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 83 | show "wf ?R" .. | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 84 | |
| 62999 | 85 | fix n :: nat | 
| 86 | assume "\<not> 100 < n" \<comment> \<open>Inner call\<close> | |
| 87 | then show "(n + 11, n) \<in> ?R" by simp | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 88 | |
| 62999 | 89 | assume inner_trm: "f91_dom (n + 11)" \<comment> \<open>Outer call\<close> | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 90 | with f91_estimate have "n + 11 < f91 (n + 11) + 11" . | 
| 62999 | 91 | with \<open>\<not> 100 < n\<close> show "(f91 (n + 11), n) \<in> ?R" by simp | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 92 | qed | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 93 | |
| 62999 | 94 | text \<open>Now trivial (even though it does not belong here):\<close> | 
| 28584 | 95 | lemma "f91 n = (if 100 < n then n - 10 else 91)" | 
| 62999 | 96 | by (induct n rule: f91.induct) auto | 
| 19568 | 97 | |
| 24585 | 98 | |
| 70468 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 99 | subsubsection \<open>Here comes Takeuchi's function\<close> | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 100 | |
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 101 | definition tak_m1 where "tak_m1 = (\<lambda>(x,y,z). if x \<le> y then 0 else 1)" | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 102 | definition tak_m2 where "tak_m2 = (\<lambda>(x,y,z). nat (Max {x, y, z} - Min {x, y, z}))"
 | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 103 | definition tak_m3 where "tak_m3 = (\<lambda>(x,y,z). nat (x - Min {x, y, z}))"
 | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 104 | |
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 105 | function tak :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int" where | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 106 | "tak x y z = (if x \<le> y then y else tak (tak (x-1) y z) (tak (y-1) z x) (tak (z-1) x y))" | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 107 | by auto | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 108 | |
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 109 | lemma tak_pcorrect: | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 110 | "tak_dom (x, y, z) \<Longrightarrow> tak x y z = (if x \<le> y then y else if y \<le> z then z else x)" | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 111 | by (induction x y z rule: tak.pinduct) (auto simp: tak.psimps) | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 112 | |
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 113 | termination | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 114 |   by (relation "tak_m1 <*mlex*> tak_m2 <*mlex*> tak_m3 <*mlex*> {}")
 | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 115 | (auto simp: mlex_iff wf_mlex tak_pcorrect tak_m1_def tak_m2_def tak_m3_def min_def max_def) | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 116 | |
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 117 | theorem tak_correct: "tak x y z = (if x \<le> y then y else if y \<le> z then z else x)" | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 118 | by (induction x y z rule: tak.induct) auto | 
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 119 | |
| 
8406a2c296e0
Added Takeuchi function to HOL-ex
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 120 | |
| 61343 | 121 | subsection \<open>More general patterns\<close> | 
| 19568 | 122 | |
| 61343 | 123 | subsubsection \<open>Overlapping patterns\<close> | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 124 | |
| 62999 | 125 | text \<open> | 
| 126 | Currently, patterns must always be compatible with each other, since | |
| 127 | no automatic splitting takes place. But the following definition of | |
| 128 | GCD is OK, although patterns overlap: | |
| 129 | \<close> | |
| 19568 | 130 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 131 | fun gcd2 :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 132 | where | 
| 19568 | 133 | "gcd2 x 0 = x" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 134 | | "gcd2 0 y = y" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 135 | | "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 | 136 | else gcd2 (x - y) (Suc y))" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 137 | |
| 19568 | 138 | thm gcd2.simps | 
| 139 | thm gcd2.induct | |
| 140 | ||
| 62999 | 141 | |
| 61343 | 142 | subsubsection \<open>Guards\<close> | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 143 | |
| 62999 | 144 | text \<open>We can reformulate the above example using guarded patterns:\<close> | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 145 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 146 | function gcd3 :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 147 | where | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 148 | "gcd3 x 0 = x" | 
| 22492 | 149 | | "gcd3 0 y = y" | 
| 63067 | 150 | | "gcd3 (Suc x) (Suc y) = gcd3 (Suc x) (y - x)" if "x < y" | 
| 151 | | "gcd3 (Suc x) (Suc y) = gcd3 (x - y) (Suc y)" if "\<not> x < y" | |
| 19922 | 152 | apply (case_tac x, case_tac a, auto) | 
| 153 | apply (case_tac ba, auto) | |
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 154 | done | 
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 155 | termination by lexicographic_order | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 156 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 157 | thm gcd3.simps | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 158 | thm gcd3.induct | 
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 159 | |
| 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 160 | |
| 61343 | 161 | text \<open>General patterns allow even strange definitions:\<close> | 
| 19782 
48c4632e2c28
HOL/Tools/function_package: imporoved handling of guards, added an example
 krauss parents: 
19770diff
changeset | 162 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 163 | function ev :: "nat \<Rightarrow> bool" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 164 | where | 
| 19568 | 165 | "ev (2 * n) = True" | 
| 22492 | 166 | | "ev (2 * n + 1) = False" | 
| 61933 | 167 | proof - \<comment> \<open>completeness is more difficult here \dots\<close> | 
| 19922 | 168 | fix P :: bool | 
| 62999 | 169 | fix x :: nat | 
| 19568 | 170 | assume c1: "\<And>n. x = 2 * n \<Longrightarrow> P" | 
| 171 | and c2: "\<And>n. x = 2 * n + 1 \<Longrightarrow> P" | |
| 172 | have divmod: "x = 2 * (x div 2) + (x mod 2)" by auto | |
| 62999 | 173 | show P | 
| 174 | proof (cases "x mod 2 = 0") | |
| 175 | case True | |
| 19568 | 176 | with divmod have "x = 2 * (x div 2)" by simp | 
| 177 | with c1 show "P" . | |
| 178 | next | |
| 62999 | 179 | case False | 
| 180 | then have "x mod 2 = 1" by simp | |
| 19568 | 181 | with divmod have "x = 2 * (x div 2) + 1" by simp | 
| 182 | with c2 show "P" . | |
| 183 | qed | |
| 62999 | 184 | qed presburger+ \<comment> \<open>solve compatibility with presburger\<close> | 
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 185 | termination by lexicographic_order | 
| 19568 | 186 | |
| 187 | thm ev.simps | |
| 188 | thm ev.induct | |
| 189 | thm ev.cases | |
| 190 | ||
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 191 | |
| 61343 | 192 | subsection \<open>Mutual Recursion\<close> | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 193 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 194 | fun evn od :: "nat \<Rightarrow> bool" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 195 | where | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 196 | "evn 0 = True" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 197 | | "od 0 = False" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 198 | | "evn (Suc n) = od n" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20270diff
changeset | 199 | | "od (Suc n) = evn n" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 200 | |
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 201 | thm evn.simps | 
| 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 202 | thm od.simps | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 203 | |
| 23817 | 204 | thm evn_od.induct | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 205 | thm evn_od.termination | 
| 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 206 | |
| 53611 | 207 | thm evn.elims | 
| 208 | thm od.elims | |
| 21240 
8e75fb38522c
Made "termination by lexicographic_order" the default for "fun" definitions.
 krauss parents: 
21051diff
changeset | 209 | |
| 62999 | 210 | |
| 61343 | 211 | subsection \<open>Definitions in local contexts\<close> | 
| 22618 | 212 | |
| 62999 | 213 | locale my_monoid = | 
| 214 | fixes opr :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 215 | and un :: "'a" | |
| 216 | assumes assoc: "opr (opr x y) z = opr x (opr y z)" | |
| 217 | and lunit: "opr un x = x" | |
| 218 | and runit: "opr x un = x" | |
| 22618 | 219 | begin | 
| 220 | ||
| 221 | fun foldR :: "'a list \<Rightarrow> 'a" | |
| 222 | where | |
| 223 | "foldR [] = un" | |
| 62999 | 224 | | "foldR (x # xs) = opr x (foldR xs)" | 
| 22618 | 225 | |
| 226 | fun foldL :: "'a list \<Rightarrow> 'a" | |
| 227 | where | |
| 228 | "foldL [] = un" | |
| 229 | | "foldL [x] = x" | |
| 62999 | 230 | | "foldL (x # y # ys) = foldL (opr x y # ys)" | 
| 22618 | 231 | |
| 232 | thm foldL.simps | |
| 233 | ||
| 234 | lemma foldR_foldL: "foldR xs = foldL xs" | |
| 62999 | 235 | by (induct xs rule: foldL.induct) (auto simp:lunit runit assoc) | 
| 22618 | 236 | |
| 237 | thm foldR_foldL | |
| 238 | ||
| 239 | end | |
| 240 | ||
| 241 | thm my_monoid.foldL.simps | |
| 242 | thm my_monoid.foldR_foldL | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 243 | |
| 62999 | 244 | |
| 61933 | 245 | subsection \<open>\<open>fun_cases\<close>\<close> | 
| 53611 | 246 | |
| 61343 | 247 | subsubsection \<open>Predecessor\<close> | 
| 53611 | 248 | |
| 62999 | 249 | fun pred :: "nat \<Rightarrow> nat" | 
| 250 | where | |
| 251 | "pred 0 = 0" | |
| 252 | | "pred (Suc n) = n" | |
| 53611 | 253 | |
| 254 | thm pred.elims | |
| 255 | ||
| 62999 | 256 | lemma | 
| 257 | assumes "pred x = y" | |
| 258 | obtains "x = 0" "y = 0" | "n" where "x = Suc n" "y = n" | |
| 259 | by (fact pred.elims[OF assms]) | |
| 260 | ||
| 53611 | 261 | |
| 61343 | 262 | text \<open>If the predecessor of a number is 0, that number must be 0 or 1.\<close> | 
| 53611 | 263 | |
| 264 | fun_cases pred0E[elim]: "pred n = 0" | |
| 265 | ||
| 266 | lemma "pred n = 0 \<Longrightarrow> n = 0 \<or> n = Suc 0" | |
| 62999 | 267 | by (erule pred0E) metis+ | 
| 53611 | 268 | |
| 62999 | 269 | text \<open> | 
| 270 | Other expressions on the right-hand side also work, but whether the | |
| 271 | generated rule is useful depends on how well the simplifier can | |
| 272 | simplify it. This example works well: | |
| 273 | \<close> | |
| 53611 | 274 | |
| 275 | fun_cases pred42E[elim]: "pred n = 42" | |
| 276 | ||
| 277 | lemma "pred n = 42 \<Longrightarrow> n = 43" | |
| 62999 | 278 | by (erule pred42E) | 
| 279 | ||
| 53611 | 280 | |
| 61343 | 281 | subsubsection \<open>List to option\<close> | 
| 53611 | 282 | |
| 62999 | 283 | fun list_to_option :: "'a list \<Rightarrow> 'a option" | 
| 284 | where | |
| 285 | "list_to_option [x] = Some x" | |
| 286 | | "list_to_option _ = None" | |
| 53611 | 287 | |
| 288 | fun_cases list_to_option_NoneE: "list_to_option xs = None" | |
| 62999 | 289 | and list_to_option_SomeE: "list_to_option xs = Some x" | 
| 53611 | 290 | |
| 291 | lemma "list_to_option xs = Some y \<Longrightarrow> xs = [y]" | |
| 62999 | 292 | by (erule list_to_option_SomeE) | 
| 293 | ||
| 53611 | 294 | |
| 61343 | 295 | subsubsection \<open>Boolean Functions\<close> | 
| 53611 | 296 | |
| 62999 | 297 | fun xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" | 
| 298 | where | |
| 299 | "xor False False = False" | |
| 300 | | "xor True True = False" | |
| 301 | | "xor _ _ = True" | |
| 53611 | 302 | |
| 303 | thm xor.elims | |
| 304 | ||
| 62999 | 305 | text \<open> | 
| 306 | \<open>fun_cases\<close> does not only recognise function equations, but also works with | |
| 307 | functions that return a boolean, e.g.: | |
| 308 | \<close> | |
| 53611 | 309 | |
| 310 | fun_cases xor_TrueE: "xor a b" and xor_FalseE: "\<not>xor a b" | |
| 311 | print_theorems | |
| 312 | ||
| 62999 | 313 | |
| 61343 | 314 | subsubsection \<open>Many parameters\<close> | 
| 53611 | 315 | |
| 62999 | 316 | fun sum4 :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 317 | where "sum4 a b c d = a + b + c + d" | |
| 53611 | 318 | |
| 319 | fun_cases sum40E: "sum4 a b c d = 0" | |
| 320 | ||
| 321 | lemma "sum4 a b c d = 0 \<Longrightarrow> a = 0" | |
| 62999 | 322 | by (erule sum40E) | 
| 53611 | 323 | |
| 40111 | 324 | |
| 61343 | 325 | subsection \<open>Partial Function Definitions\<close> | 
| 40111 | 326 | |
| 61343 | 327 | text \<open>Partial functions in the option monad:\<close> | 
| 40111 | 328 | |
| 329 | partial_function (option) | |
| 330 | collatz :: "nat \<Rightarrow> nat list option" | |
| 331 | where | |
| 332 | "collatz n = | |
| 62999 | 333 | (if n \<le> 1 then Some [n] | 
| 334 | else if even n | |
| 335 |        then do { ns \<leftarrow> collatz (n div 2); Some (n # ns) }
 | |
| 336 |        else do { ns \<leftarrow> collatz (3 * n + 1);  Some (n # ns)})"
 | |
| 40111 | 337 | |
| 40169 
11ea439d947f
declare recursive equation as ".simps", in accordance with other packages
 krauss parents: 
40111diff
changeset | 338 | declare collatz.simps[code] | 
| 40111 | 339 | value "collatz 23" | 
| 340 | ||
| 341 | ||
| 61343 | 342 | text \<open>Tail-recursive functions:\<close> | 
| 40111 | 343 | |
| 344 | partial_function (tailrec) fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
 | |
| 345 | where | |
| 346 | "fixpoint f x = (if f x = x then x else fixpoint f (f x))" | |
| 347 | ||
| 348 | ||
| 61343 | 349 | subsection \<open>Regression tests\<close> | 
| 22726 | 350 | |
| 62999 | 351 | text \<open> | 
| 352 | The following examples mainly serve as tests for the | |
| 353 | function package. | |
| 354 | \<close> | |
| 22726 | 355 | |
| 356 | fun listlen :: "'a list \<Rightarrow> nat" | |
| 357 | where | |
| 358 | "listlen [] = 0" | |
| 359 | | "listlen (x#xs) = Suc (listlen xs)" | |
| 360 | ||
| 361 | ||
| 62999 | 362 | subsubsection \<open>Context recursion\<close> | 
| 363 | ||
| 364 | fun f :: "nat \<Rightarrow> nat" | |
| 22726 | 365 | where | 
| 366 | zero: "f 0 = 0" | |
| 367 | | succ: "f (Suc n) = (if f n = 0 then 0 else f n)" | |
| 368 | ||
| 369 | ||
| 62999 | 370 | subsubsection \<open>A combination of context and nested recursion\<close> | 
| 371 | ||
| 22726 | 372 | function h :: "nat \<Rightarrow> nat" | 
| 373 | where | |
| 374 | "h 0 = 0" | |
| 375 | | "h (Suc n) = (if h n = 0 then h (h n) else h n)" | |
| 62999 | 376 | by pat_completeness auto | 
| 22726 | 377 | |
| 378 | ||
| 62999 | 379 | subsubsection \<open>Context, but no recursive call\<close> | 
| 380 | ||
| 22726 | 381 | fun i :: "nat \<Rightarrow> nat" | 
| 382 | where | |
| 383 | "i 0 = 0" | |
| 384 | | "i (Suc n) = (if n = 0 then 0 else i n)" | |
| 385 | ||
| 386 | ||
| 62999 | 387 | subsubsection \<open>Tupled nested recursion\<close> | 
| 388 | ||
| 22726 | 389 | fun fa :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 390 | where | |
| 391 | "fa 0 y = 0" | |
| 392 | | "fa (Suc n) y = (if fa n y = 0 then 0 else fa n y)" | |
| 393 | ||
| 62999 | 394 | |
| 395 | subsubsection \<open>Let\<close> | |
| 396 | ||
| 22726 | 397 | fun j :: "nat \<Rightarrow> nat" | 
| 398 | where | |
| 399 | "j 0 = 0" | |
| 62999 | 400 | | "j (Suc n) = (let u = n in Suc (j u))" | 
| 22726 | 401 | |
| 402 | ||
| 62999 | 403 | text \<open>There were some problems with fresh names \dots\<close> | 
| 22726 | 404 | function k :: "nat \<Rightarrow> nat" | 
| 405 | where | |
| 406 | "k x = (let a = x; b = x in k x)" | |
| 407 | by pat_completeness auto | |
| 408 | ||
| 409 | ||
| 410 | function f2 :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)" | |
| 411 | where | |
| 412 | "f2 p = (let (x,y) = p in f2 (y,x))" | |
| 413 | by pat_completeness auto | |
| 414 | ||
| 415 | ||
| 62999 | 416 | subsubsection \<open>Abbreviations\<close> | 
| 417 | ||
| 22726 | 418 | fun f3 :: "'a set \<Rightarrow> bool" | 
| 419 | where | |
| 420 | "f3 x = finite x" | |
| 421 | ||
| 422 | ||
| 62999 | 423 | subsubsection \<open>Simple Higher-Order Recursion\<close> | 
| 424 | ||
| 425 | datatype 'a tree = Leaf 'a | Branch "'a tree list" | |
| 23117 
e2744f32641e
updated examples to include an instance of (lexicographic_order simp:...)
 krauss parents: 
22726diff
changeset | 426 | |
| 36269 | 427 | fun treemap :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a tree \<Rightarrow> 'a tree"
 | 
| 22726 | 428 | where | 
| 429 | "treemap fn (Leaf n) = (Leaf (fn n))" | |
| 430 | | "treemap fn (Branch l) = (Branch (map (treemap fn) l))" | |
| 431 | ||
| 432 | fun tinc :: "nat tree \<Rightarrow> nat tree" | |
| 433 | where | |
| 434 | "tinc (Leaf n) = Leaf (Suc n)" | |
| 435 | | "tinc (Branch l) = Branch (map tinc l)" | |
| 436 | ||
| 36270 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 437 | fun testcase :: "'a tree \<Rightarrow> 'a list" | 
| 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 438 | where | 
| 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 439 | "testcase (Leaf a) = [a]" | 
| 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 440 | | "testcase (Branch x) = | 
| 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 441 | (let xs = concat (map testcase x); | 
| 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 442 | ys = concat (map testcase x) in | 
| 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 443 | xs @ ys)" | 
| 
fd95c0514623
tolerate eta-variants in f_graph.cases (from inductive package); added test case;
 krauss parents: 
36269diff
changeset | 444 | |
| 22726 | 445 | |
| 62999 | 446 | subsubsection \<open>Pattern matching on records\<close> | 
| 447 | ||
| 22726 | 448 | record point = | 
| 449 | Xcoord :: int | |
| 450 | Ycoord :: int | |
| 451 | ||
| 452 | function swp :: "point \<Rightarrow> point" | |
| 453 | where | |
| 454 | "swp \<lparr> Xcoord = x, Ycoord = y \<rparr> = \<lparr> Xcoord = y, Ycoord = x \<rparr>" | |
| 455 | proof - | |
| 456 | fix P x | |
| 457 | assume "\<And>xa y. x = \<lparr>Xcoord = xa, Ycoord = y\<rparr> \<Longrightarrow> P" | |
| 62999 | 458 | then show P by (cases x) | 
| 22726 | 459 | qed auto | 
| 460 | termination by rule auto | |
| 461 | ||
| 462 | ||
| 62999 | 463 | subsubsection \<open>The diagonal function\<close> | 
| 464 | ||
| 22726 | 465 | fun diag :: "bool \<Rightarrow> bool \<Rightarrow> bool \<Rightarrow> nat" | 
| 466 | where | |
| 467 | "diag x True False = 1" | |
| 468 | | "diag False y True = 2" | |
| 469 | | "diag True False z = 3" | |
| 470 | | "diag True True True = 4" | |
| 471 | | "diag False False False = 5" | |
| 472 | ||
| 473 | ||
| 62999 | 474 | subsubsection \<open>Many equations (quadratic blowup)\<close> | 
| 475 | ||
| 476 | datatype DT = | |
| 22726 | 477 | A | B | C | D | E | F | G | H | I | J | K | L | M | N | P | 
| 478 | | Q | R | S | T | U | V | |
| 479 | ||
| 480 | fun big :: "DT \<Rightarrow> nat" | |
| 481 | where | |
| 62999 | 482 | "big A = 0" | 
| 483 | | "big B = 0" | |
| 484 | | "big C = 0" | |
| 485 | | "big D = 0" | |
| 486 | | "big E = 0" | |
| 487 | | "big F = 0" | |
| 488 | | "big G = 0" | |
| 489 | | "big H = 0" | |
| 490 | | "big I = 0" | |
| 491 | | "big J = 0" | |
| 492 | | "big K = 0" | |
| 493 | | "big L = 0" | |
| 494 | | "big M = 0" | |
| 495 | | "big N = 0" | |
| 496 | | "big P = 0" | |
| 497 | | "big Q = 0" | |
| 498 | | "big R = 0" | |
| 499 | | "big S = 0" | |
| 500 | | "big T = 0" | |
| 501 | | "big U = 0" | |
| 22726 | 502 | | "big V = 0" | 
| 503 | ||
| 504 | ||
| 62999 | 505 | subsubsection \<open>Automatic pattern splitting\<close> | 
| 506 | ||
| 507 | fun f4 :: "nat \<Rightarrow> nat \<Rightarrow> bool" | |
| 22726 | 508 | where | 
| 509 | "f4 0 0 = True" | |
| 25170 | 510 | | "f4 _ _ = False" | 
| 22726 | 511 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19736diff
changeset | 512 | |
| 63014 | 513 | subsubsection \<open>Polymorphic partial-function\<close> | 
| 62999 | 514 | |
| 45008 
8b74cfea913a
match types when applying mono_thm -- previous export generalizes type variables;
 krauss parents: 
41817diff
changeset | 515 | partial_function (option) f5 :: "'a list \<Rightarrow> 'a option" | 
| 
8b74cfea913a
match types when applying mono_thm -- previous export generalizes type variables;
 krauss parents: 
41817diff
changeset | 516 | where | 
| 
8b74cfea913a
match types when applying mono_thm -- previous export generalizes type variables;
 krauss parents: 
41817diff
changeset | 517 | "f5 x = f5 x" | 
| 
8b74cfea913a
match types when applying mono_thm -- previous export generalizes type variables;
 krauss parents: 
41817diff
changeset | 518 | |
| 19736 | 519 | end |