author | blanchet |
Fri, 27 May 2011 10:30:08 +0200 | |
changeset 43016 | 42330f25142c |
parent 37595 | 9591362629e3 |
child 44013 | 5cfc1c36ae97 |
permissions | -rw-r--r-- |
6481 | 1 |
(* Title: HOL/ex/Recdefs.thy |
5124 | 2 |
Author: Konrad Slind and Lawrence C Paulson |
3 |
Copyright 1996 University of Cambridge |
|
4 |
||
5 |
Examples of recdef definitions. Most, but not all, are handled automatically. |
|
6 |
*) |
|
7 |
||
11024 | 8 |
header {* Examples of recdef definitions *} |
9 |
||
16417 | 10 |
theory Recdefs imports Main begin |
5124 | 11 |
|
12 |
consts fact :: "nat => nat" |
|
11024 | 13 |
recdef fact less_than |
14 |
"fact x = (if x = 0 then 1 else x * fact (x - 1))" |
|
5124 | 15 |
|
16 |
consts Fact :: "nat => nat" |
|
11024 | 17 |
recdef Fact less_than |
18 |
"Fact 0 = 1" |
|
19 |
"Fact (Suc x) = Fact x * Suc x" |
|
5124 | 20 |
|
14953 | 21 |
consts fib :: "int => int" |
22 |
recdef fib "measure nat" |
|
23 |
eqn: "fib n = (if n < 1 then 0 |
|
24 |
else if n=1 then 1 |
|
25 |
else fib(n - 2) + fib(n - 1))"; |
|
26 |
||
27 |
lemma "fib 7 = 13" |
|
28 |
by simp |
|
29 |
||
30 |
||
5124 | 31 |
consts map2 :: "('a => 'b => 'c) * 'a list * 'b list => 'c list" |
11024 | 32 |
recdef map2 "measure(\<lambda>(f, l1, l2). size l1)" |
33 |
"map2 (f, [], []) = []" |
|
34 |
"map2 (f, h # t, []) = []" |
|
35 |
"map2 (f, h1 # t1, h2 # t2) = f h1 h2 # map2 (f, t1, t2)" |
|
5124 | 36 |
|
11024 | 37 |
consts finiteRchain :: "('a => 'a => bool) * 'a list => bool" |
38 |
recdef finiteRchain "measure (\<lambda>(R, l). size l)" |
|
39 |
"finiteRchain(R, []) = True" |
|
40 |
"finiteRchain(R, [x]) = True" |
|
41 |
"finiteRchain(R, x # y # rst) = (R x y \<and> finiteRchain (R, y # rst))" |
|
5124 | 42 |
|
11024 | 43 |
text {* Not handled automatically: too complicated. *} |
5124 | 44 |
consts variant :: "nat * nat list => nat" |
14244
f58598341d30
InductiveInvariant_examples illustrates advanced recursive function definitions
paulson
parents:
12023
diff
changeset
|
45 |
recdef (permissive) variant "measure (\<lambda>(n,ns). size (filter (\<lambda>y. n \<le> y) ns))" |
37595
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
haftmann
parents:
37456
diff
changeset
|
46 |
"variant (x, L) = (if x \<in> set L then variant (Suc x, L) else x)" |
5124 | 47 |
|
48 |
consts gcd :: "nat * nat => nat" |
|
11024 | 49 |
recdef gcd "measure (\<lambda>(x, y). x + y)" |
50 |
"gcd (0, y) = y" |
|
51 |
"gcd (Suc x, 0) = Suc x" |
|
52 |
"gcd (Suc x, Suc y) = |
|
53 |
(if y \<le> x then gcd (x - y, Suc y) else gcd (Suc x, y - x))" |
|
54 |
||
55 |
||
56 |
text {* |
|
57 |
\medskip The silly @{term g} function: example of nested recursion. |
|
58 |
Not handled automatically. In fact, @{term g} is the zero constant |
|
59 |
function. |
|
60 |
*} |
|
5124 | 61 |
|
11024 | 62 |
consts g :: "nat => nat" |
11626 | 63 |
recdef (permissive) g less_than |
11024 | 64 |
"g 0 = 0" |
65 |
"g (Suc x) = g (g x)" |
|
66 |
||
67 |
lemma g_terminates: "g x < Suc x" |
|
68 |
apply (induct x rule: g.induct) |
|
69 |
apply (auto simp add: g.simps) |
|
70 |
done |
|
71 |
||
72 |
lemma g_zero: "g x = 0" |
|
73 |
apply (induct x rule: g.induct) |
|
74 |
apply (simp_all add: g.simps g_terminates) |
|
75 |
done |
|
76 |
||
5124 | 77 |
|
78 |
consts Div :: "nat * nat => nat * nat" |
|
11024 | 79 |
recdef Div "measure fst" |
80 |
"Div (0, x) = (0, 0)" |
|
81 |
"Div (Suc x, y) = |
|
82 |
(let (q, r) = Div (x, y) |
|
83 |
in if y \<le> Suc r then (Suc q, 0) else (q, Suc r))" |
|
84 |
||
85 |
text {* |
|
86 |
\medskip Not handled automatically. Should be the predecessor |
|
87 |
function, but there is an unnecessary "looping" recursive call in |
|
12023 | 88 |
@{text "k 1"}. |
11024 | 89 |
*} |
5124 | 90 |
|
11024 | 91 |
consts k :: "nat => nat" |
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11626
diff
changeset
|
92 |
|
11626 | 93 |
recdef (permissive) k less_than |
11024 | 94 |
"k 0 = 0" |
95 |
"k (Suc n) = |
|
96 |
(let x = k 1 |
|
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
11626
diff
changeset
|
97 |
in if False then k (Suc 1) else n)" |
11024 | 98 |
|
99 |
consts part :: "('a => bool) * 'a list * 'a list * 'a list => 'a list * 'a list" |
|
100 |
recdef part "measure (\<lambda>(P, l, l1, l2). size l)" |
|
101 |
"part (P, [], l1, l2) = (l1, l2)" |
|
102 |
"part (P, h # rst, l1, l2) = |
|
103 |
(if P h then part (P, rst, h # l1, l2) |
|
104 |
else part (P, rst, l1, h # l2))" |
|
5124 | 105 |
|
11024 | 106 |
consts fqsort :: "('a => 'a => bool) * 'a list => 'a list" |
11626 | 107 |
recdef (permissive) fqsort "measure (size o snd)" |
11024 | 108 |
"fqsort (ord, []) = []" |
109 |
"fqsort (ord, x # rst) = |
|
110 |
(let (less, more) = part ((\<lambda>y. ord y x), rst, ([], [])) |
|
111 |
in fqsort (ord, less) @ [x] @ fqsort (ord, more))" |
|
112 |
||
113 |
text {* |
|
114 |
\medskip Silly example which demonstrates the occasional need for |
|
115 |
additional congruence rules (here: @{thm [source] map_cong}). If |
|
116 |
the congruence rule is removed, an unprovable termination condition |
|
117 |
is generated! Termination not proved automatically. TFL requires |
|
118 |
@{term [source] "\<lambda>x. mapf x"} instead of @{term [source] mapf}. |
|
119 |
*} |
|
5124 | 120 |
|
11024 | 121 |
consts mapf :: "nat => nat list" |
37456 | 122 |
recdef mapf "measure (\<lambda>m. m)" |
11024 | 123 |
"mapf 0 = []" |
124 |
"mapf (Suc n) = concat (map (\<lambda>x. mapf x) (replicate n n))" |
|
125 |
(hints cong: map_cong) |
|
5124 | 126 |
|
37456 | 127 |
(* This used to be an example where recdef_tc etc is needed, |
128 |
but now termination is proved outright |
|
11024 | 129 |
recdef_tc mapf_tc: mapf |
130 |
apply (rule allI) |
|
131 |
apply (case_tac "n = 0") |
|
132 |
apply simp_all |
|
133 |
done |
|
134 |
||
135 |
text {* Removing the termination condition from the generated thms: *} |
|
136 |
||
137 |
lemma "mapf (Suc n) = concat (map mapf (replicate n n))" |
|
138 |
apply (simp add: mapf.simps mapf_tc) |
|
139 |
done |
|
140 |
||
141 |
lemmas mapf_induct = mapf.induct [OF mapf_tc] |
|
37456 | 142 |
*) |
5124 | 143 |
|
144 |
end |