src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy
changeset 38792 970508a5119f
parent 38791 4a4be1be0fae
child 38949 1afa9e89c885
equal deleted inserted replaced
38791:4a4be1be0fae 38792:970508a5119f
     2 imports Code_Prolog
     2 imports Code_Prolog
     3 begin
     3 begin
     4 
     4 
     5 section {* Example append *}
     5 section {* Example append *}
     6 
     6 
       
     7 
     7 inductive append
     8 inductive append
     8 where
     9 where
     9   "append [] ys ys"
    10   "append [] ys ys"
    10 | "append xs ys zs ==> append (x # xs) ys (x # zs)"
    11 | "append xs ys zs ==> append (x # xs) ys (x # zs)"
    11 
    12 
       
    13 ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
       
    14 
       
    15 values "{(x, y, z). append x y z}"
       
    16 
    12 values 3 "{(x, y, z). append x y z}"
    17 values 3 "{(x, y, z). append x y z}"
       
    18 
       
    19 ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.YAP} *}
       
    20 
       
    21 values "{(x, y, z). append x y z}"
       
    22 
       
    23 ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
    13 
    24 
    14 
    25 
    15 section {* Example queens *}
    26 section {* Example queens *}
    16 
    27 
    17 inductive nodiag :: "int => int => int list => bool"
    28 inductive nodiag :: "int => int => int list => bool"
   170 
   181 
   171 inductive notB :: "abc => bool"
   182 inductive notB :: "abc => bool"
   172 where
   183 where
   173   "y \<noteq> B \<Longrightarrow> notB y"
   184   "y \<noteq> B \<Longrightarrow> notB y"
   174 
   185 
   175 ML {* Code_Prolog.options := {ensure_groundness = true, limited_types = []} *}
   186 ML {* Code_Prolog.options := {ensure_groundness = true, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
   176 
   187 
   177 values 2 "{y. notB y}"
   188 values 2 "{y. notB y}"
   178 
   189 
   179 inductive notAB :: "abc * abc \<Rightarrow> bool"
   190 inductive notAB :: "abc * abc \<Rightarrow> bool"
   180 where
   191 where