--- a/src/HOL/ex/Predicate_Compile.thy Sun Mar 29 19:48:35 2009 +0200
+++ b/src/HOL/ex/Predicate_Compile.thy Tue Mar 31 11:04:05 2009 +0200
@@ -1,5 +1,5 @@
theory Predicate_Compile
-imports Complex_Main Lattice_Syntax
+imports Complex_Main Code_Index Lattice_Syntax
uses "predicate_compile.ML"
begin
@@ -12,10 +12,27 @@
| "next yield (Predicate.Join P xq) = (case yield P
of None \<Rightarrow> next yield xq | Some (x, Q) \<Rightarrow> Some (x, Predicate.Seq (\<lambda>_. Predicate.Join Q xq)))"
-primrec pull :: "('a Predicate.pred \<Rightarrow> ('a \<times> 'a Predicate.pred) option)
- \<Rightarrow> nat \<Rightarrow> 'a Predicate.pred \<Rightarrow> 'a list \<times> 'a Predicate.pred" where
- "pull yield 0 P = ([], \<bottom>)"
- | "pull yield (Suc n) P = (case yield P
- of None \<Rightarrow> ([], \<bottom>) | Some (x, Q) \<Rightarrow> let (xs, R) = pull yield n Q in (x # xs, R))"
+ML {*
+let
+ fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ())
+in
+ yield @{code "\<bottom> :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*)
+end
+*}
+
+fun pull :: "('a Predicate.pred \<Rightarrow> ('a \<times> 'a Predicate.pred) option)
+ \<Rightarrow> index \<Rightarrow> 'a Predicate.pred \<Rightarrow> 'a list \<times> 'a Predicate.pred" where
+ "pull yield k P = (if k = 0 then ([], \<bottom>)
+ else case yield P of None \<Rightarrow> ([], \<bottom>) | Some (x, Q) \<Rightarrow> let (xs, R) = pull yield (k - 1) Q in (x # xs, R))"
+
+ML {*
+let
+ fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ())
+ fun yieldn k = @{code pull} yield k
+in
+ yieldn 0 (*replace with number of elements to retrieve*)
+ @{code "\<bottom> :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*)
+end
+*}
end
\ No newline at end of file