author | bulwahn |
Mon, 11 May 2009 09:39:53 +0200 | |
changeset 31107 | 657386d94f14 |
parent 31106 | 9a1178204dc0 |
child 31108 | 0ce5f53fc65d |
permissions | -rw-r--r-- |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
1 |
theory Predicate_Compile |
30810 | 2 |
imports Complex_Main Code_Index Lattice_Syntax |
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
3 |
uses "predicate_compile.ML" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
4 |
begin |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
5 |
|
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
6 |
setup {* Predicate_Compile.setup *} |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
7 |
|
31106 | 8 |
ML {* |
9 |
OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate" |
|
10 |
OuterKeyword.thy_goal (OuterParse.term_group >> Predicate_Compile.code_pred_cmd) |
|
11 |
*} |
|
12 |
||
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
13 |
primrec "next" :: "('a Predicate.pred \<Rightarrow> ('a \<times> 'a Predicate.pred) option) |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
14 |
\<Rightarrow> 'a Predicate.seq \<Rightarrow> ('a \<times> 'a Predicate.pred) option" where |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
15 |
"next yield Predicate.Empty = None" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
16 |
| "next yield (Predicate.Insert x P) = Some (x, P)" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
17 |
| "next yield (Predicate.Join P xq) = (case yield P |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
18 |
of None \<Rightarrow> next yield xq | Some (x, Q) \<Rightarrow> Some (x, Predicate.Seq (\<lambda>_. Predicate.Join Q xq)))" |
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
19 |
|
30810 | 20 |
ML {* |
21 |
let |
|
22 |
fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ()) |
|
23 |
in |
|
24 |
yield @{code "\<bottom> :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*) |
|
25 |
end |
|
26 |
*} |
|
27 |
||
30812 | 28 |
fun anamorph :: "('b \<Rightarrow> ('a \<times> 'b) option) \<Rightarrow> index \<Rightarrow> 'b \<Rightarrow> 'a list \<times> 'b" where |
29 |
"anamorph f k x = (if k = 0 then ([], x) |
|
30 |
else case f x of None \<Rightarrow> ([], x) | Some (v, y) \<Rightarrow> let (vs, z) = anamorph f (k - 1) y in (v # vs, z))" |
|
30810 | 31 |
|
32 |
ML {* |
|
33 |
let |
|
34 |
fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ()) |
|
30812 | 35 |
fun yieldn k = @{code anamorph} yield k |
30810 | 36 |
in |
37 |
yieldn 0 (*replace with number of elements to retrieve*) |
|
38 |
@{code "\<bottom> :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*) |
|
39 |
end |
|
40 |
*} |
|
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
41 |
|
31106 | 42 |
section {* Example for user interface *} |
43 |
||
44 |
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" |
|
45 |
where |
|
46 |
"append [] ys ys" |
|
47 |
| "append xs' ys zs' \<Longrightarrow> append (x#xs') ys (x#zs')" |
|
48 |
||
49 |
code_pred append |
|
50 |
using assms by (rule append.cases) |
|
51 |
||
31107 | 52 |
thm append_codegen |
31106 | 53 |
thm append_cases |
54 |
||
55 |
||
30374
7311a1546d85
added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff
changeset
|
56 |
end |