src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy
changeset 38963 b5d126d7be4b
parent 38950 62578950e748
child 39189 d183bf90dabd
equal deleted inserted replaced
38962:3917c2acaec4 38963:b5d126d7be4b
     8 inductive append
     8 inductive append
     9 where
     9 where
    10   "append [] ys ys"
    10   "append [] ys ys"
    11 | "append xs ys zs ==> append (x # xs) ys (x # zs)"
    11 | "append xs ys zs ==> append (x # xs) ys (x # zs)"
    12 
    12 
    13 setup {* Code_Prolog.map_code_options (K {ensure_groundness = false,
    13 setup {* Code_Prolog.map_code_options (K
    14    limited_types = [], limited_predicates = [], replacing = [],
    14   {ensure_groundness = false,
       
    15    limited_types = [],
       
    16    limited_predicates = [],
       
    17    replacing = [],
       
    18    manual_reorder = [],
    15    prolog_system = Code_Prolog.SWI_PROLOG}) *}
    19    prolog_system = Code_Prolog.SWI_PROLOG}) *}
    16 
    20 
    17 values "{(x, y, z). append x y z}"
    21 values "{(x, y, z). append x y z}"
    18 
    22 
    19 values 3 "{(x, y, z). append x y z}"
    23 values 3 "{(x, y, z). append x y z}"
    20 
    24 
    21 setup {* Code_Prolog.map_code_options (K
    25 setup {* Code_Prolog.map_code_options (K
    22   {ensure_groundness = false,
    26   {ensure_groundness = false,
    23    limited_types = [], limited_predicates = [], replacing = [],
    27    limited_types = [],
       
    28    limited_predicates = [],
       
    29    replacing = [],
       
    30    manual_reorder = [],
    24    prolog_system = Code_Prolog.YAP}) *}
    31    prolog_system = Code_Prolog.YAP}) *}
    25 
    32 
    26 values "{(x, y, z). append x y z}"
    33 values "{(x, y, z). append x y z}"
    27 
    34 
    28 setup {* Code_Prolog.map_code_options (K
    35 setup {* Code_Prolog.map_code_options (K
    29   {ensure_groundness = false,
    36   {ensure_groundness = false,
    30    limited_types = [], limited_predicates = [], replacing = [],
    37    limited_types = [],
       
    38    limited_predicates = [],
       
    39    replacing = [],
       
    40    manual_reorder = [],
    31    prolog_system = Code_Prolog.SWI_PROLOG}) *}
    41    prolog_system = Code_Prolog.SWI_PROLOG}) *}
    32 
    42 
    33 
    43 
    34 section {* Example queens *}
    44 section {* Example queens *}
    35 
    45 
   189 
   199 
   190 inductive notB :: "abc => bool"
   200 inductive notB :: "abc => bool"
   191 where
   201 where
   192   "y \<noteq> B \<Longrightarrow> notB y"
   202   "y \<noteq> B \<Longrightarrow> notB y"
   193 
   203 
   194 setup {* Code_Prolog.map_code_options (K {ensure_groundness = true,
   204 setup {* Code_Prolog.map_code_options (K
   195   limited_types = [],
   205   {ensure_groundness = true,
   196   limited_predicates = [],
   206    limited_types = [],
   197   replacing = [], 
   207    limited_predicates = [],
   198   prolog_system = Code_Prolog.SWI_PROLOG}) *}
   208    replacing = [],
       
   209    manual_reorder = [], 
       
   210    prolog_system = Code_Prolog.SWI_PROLOG}) *}
   199 
   211 
   200 values 2 "{y. notB y}"
   212 values 2 "{y. notB y}"
   201 
   213 
   202 inductive notAB :: "abc * abc \<Rightarrow> bool"
   214 inductive notAB :: "abc * abc \<Rightarrow> bool"
   203 where
   215 where