src/HOL/ex/Predicate_Compile.thy
author bulwahn
Mon, 11 May 2009 09:39:53 +0200
changeset 31107 657386d94f14
parent 31106 9a1178204dc0
child 31108 0ce5f53fc65d
permissions -rw-r--r--
fixed code_pred command
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
     1
theory Predicate_Compile
30810
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
     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
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
     8
ML {*
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
     9
  OuterSyntax.local_theory_to_proof "code_pred" "sets up goal for cases rule from given introduction rules and compiles predicate"
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    10
  OuterKeyword.thy_goal (OuterParse.term_group >> Predicate_Compile.code_pred_cmd)
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    11
*}
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    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
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    20
ML {*
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    21
let
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    22
  fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ())
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    23
in
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    24
  yield @{code "\<bottom> :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*)
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    25
end
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    26
*}
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    27
30812
7d02340f095d generalized pull to anamorph
haftmann
parents: 30810
diff changeset
    28
fun anamorph :: "('b \<Rightarrow> ('a \<times> 'b) option) \<Rightarrow> index \<Rightarrow> 'b \<Rightarrow> 'a list \<times> 'b" where
7d02340f095d generalized pull to anamorph
haftmann
parents: 30810
diff changeset
    29
  "anamorph f k x = (if k = 0 then ([], x)
7d02340f095d generalized pull to anamorph
haftmann
parents: 30810
diff changeset
    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
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    31
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    32
ML {*
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    33
let
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    34
  fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ())
30812
7d02340f095d generalized pull to anamorph
haftmann
parents: 30810
diff changeset
    35
  fun yieldn k = @{code anamorph} yield k
30810
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    36
in
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    37
  yieldn 0 (*replace with number of elements to retrieve*)
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    38
    @{code "\<bottom> :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*)
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    39
end
83642621425a ML snippets for experimental evaluation
haftmann
parents: 30374
diff changeset
    40
*}
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
    41
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    42
section {* Example for user interface *}
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    43
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    44
inductive append ::  "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    45
where
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    46
  "append [] ys ys"
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    47
| "append xs' ys zs' \<Longrightarrow> append (x#xs') ys (x#zs')"
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    48
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    49
code_pred append
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    50
  using assms by (rule append.cases)
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    51
31107
657386d94f14 fixed code_pred command
bulwahn
parents: 31106
diff changeset
    52
thm append_codegen
31106
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    53
thm append_cases
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    54
9a1178204dc0 Added pred_code command
bulwahn
parents: 30812
diff changeset
    55
30374
7311a1546d85 added predicate compiler, as formally checked prototype, not as user package
haftmann
parents:
diff changeset
    56
end