adapting example files to latest changes
authorbulwahn
Tue, 31 Aug 2010 08:00:52 +0200
changeset 38949 1afa9e89c885
parent 38948 c4e6afaa8dcd
child 38950 62578950e748
adapting example files to latest changes
src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy
src/HOL/Predicate_Compile_Examples/Hotel_Example.thy
--- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Tue Aug 31 08:00:51 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Tue Aug 31 08:00:52 2010 +0200
@@ -10,17 +10,26 @@
   "append [] ys ys"
 | "append xs ys zs ==> append (x # xs) ys (x # zs)"
 
-ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
+ML {* Code_Prolog.options :=
+  {ensure_groundness = false,
+   limited_types = [], limited_predicates = [], replacing = [],
+   prolog_system = Code_Prolog.SWI_PROLOG} *}
 
 values "{(x, y, z). append x y z}"
 
 values 3 "{(x, y, z). append x y z}"
 
-ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.YAP} *}
+ML {* Code_Prolog.options :=
+  {ensure_groundness = false,
+   limited_types = [], limited_predicates = [], replacing = [],
+   prolog_system = Code_Prolog.YAP} *}
 
 values "{(x, y, z). append x y z}"
 
-ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
+ML {* Code_Prolog.options :=
+  {ensure_groundness = false,
+   limited_types = [], limited_predicates = [], replacing = [],
+   prolog_system = Code_Prolog.SWI_PROLOG} *}
 
 
 section {* Example queens *}
@@ -183,7 +192,11 @@
 where
   "y \<noteq> B \<Longrightarrow> notB y"
 
-ML {* Code_Prolog.options := {ensure_groundness = true, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
+ML {* Code_Prolog.options := {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [],
+  replacing = [], 
+  prolog_system = Code_Prolog.SWI_PROLOG} *}
 
 values 2 "{y. notB y}"
 
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy	Tue Aug 31 08:00:51 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy	Tue Aug 31 08:00:52 2010 +0200
@@ -84,17 +84,21 @@
 lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
 by (auto simp add: Diff_iff[unfolded mem_def] expand_fun_eq intro!: eq_reflection)
 
-ML {* Code_Prolog.options := {ensure_groundness = true} *}
+ML {* Code_Prolog.options :=
+  {ensure_groundness = true,
+  limited_types = [],
+  limited_predicates = [],
+  replacing = [],
+  prolog_system = Code_Prolog.SWI_PROLOG} *}
 
 values 40 "{s. hotel s}"
 
 
 setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
-ML {* set Code_Prolog.trace *}
 
 lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
 quickcheck[generator = code, iterations = 100000, report]
-quickcheck[generator = prolog, iterations = 1]
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
 oops