src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy
changeset 38792 970508a5119f
parent 38791 4a4be1be0fae
child 38949 1afa9e89c885
--- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Thu Aug 26 14:48:48 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Fri Aug 27 09:34:06 2010 +0200
@@ -4,13 +4,24 @@
 
 section {* Example append *}
 
+
 inductive append
 where
   "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} *}
+
+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} *}
+
+values "{(x, y, z). append x y z}"
+
+ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
+
 
 section {* Example queens *}
 
@@ -172,7 +183,7 @@
 where
   "y \<noteq> B \<Longrightarrow> notB y"
 
-ML {* Code_Prolog.options := {ensure_groundness = true, limited_types = []} *}
+ML {* Code_Prolog.options := {ensure_groundness = true, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
 
 values 2 "{y. notB y}"