diff -r 143f58bb34f9 -r 0909deb8059b src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy --- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Thu May 26 16:57:14 2016 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Thu May 26 17:51:22 2016 +0200 @@ -2,7 +2,7 @@ imports "~~/src/HOL/Library/Code_Prolog" begin -section {* Example append *} +section \Example append\ inductive append @@ -10,13 +10,13 @@ "append [] ys ys" | "append xs ys zs ==> append (x # xs) ys (x # zs)" -setup {* Code_Prolog.map_code_options (K +setup \Code_Prolog.map_code_options (K {ensure_groundness = false, limit_globally = NONE, limited_types = [], limited_predicates = [], replacing = [], - manual_reorder = []}) *} + manual_reorder = []})\ values_prolog "{(x, y, z). append x y z}" @@ -24,73 +24,73 @@ values_prolog 3 "{(x, y, z). append x y z}" -setup {* Code_Prolog.map_code_options (K +setup \Code_Prolog.map_code_options (K {ensure_groundness = false, limit_globally = NONE, limited_types = [], limited_predicates = [], replacing = [], - manual_reorder = []}) *} + manual_reorder = []})\ values_prolog "{(x, y, z). append x y z}" -setup {* Code_Prolog.map_code_options (K +setup \Code_Prolog.map_code_options (K {ensure_groundness = false, limit_globally = NONE, limited_types = [], limited_predicates = [], replacing = [], - manual_reorder = []}) *} + manual_reorder = []})\ -section {* Example queens *} +section \Example queens\ inductive nodiag :: "int => int => int list => bool" where "nodiag B D []" | "D \ N - B ==> D \ B - N ==> Da = D + 1 ==> nodiag B Da L ==> nodiag B D (N # L)" -text {* +text \ qdelete(A, [A|L], L). qdelete(X, [A|Z], [A|R]) :- qdelete(X, Z, R). -*} +\ inductive qdelete :: "int => int list => int list => bool" where "qdelete A (A # L) L" | "qdelete X Z R ==> qdelete X (A # Z) (A # R)" -text {* +text \ qperm([], []). qperm([X|Y], K) :- qdelete(U, [X|Y], Z), K = [U|V], qperm(Z, V). -*} +\ inductive qperm :: "int list => int list => bool" where "qperm [] []" | "qdelete U (X # Y) Z ==> qperm Z V ==> qperm (X # Y) (U # V)" -text {* +text \ safe([]). safe([N|L]) :- nodiag(N, 1, L), safe(L). -*} +\ inductive safe :: "int list => bool" where "safe []" | "nodiag N 1 L ==> safe L ==> safe (N # L)" -text {* +text \ queen(Data, Out) :- qperm(Data, Out), safe(Out) -*} +\ inductive queen :: "int list => int list => bool" where @@ -103,14 +103,14 @@ values_prolog 10 "{ys. queen_9 ys}" -section {* Example symbolic derivation *} +section \Example symbolic derivation\ hide_const Pow datatype expr = Log expr | Mult expr expr | Div expr expr | x | Num int | Plus expr expr | Minus expr expr | Uminus expr | Pow expr int | Exp expr -text {* +text \ d(U + V, X, DU + DV) :- cut, @@ -145,7 +145,7 @@ cut. d(num(_), _, num(0)). -*} +\ inductive d :: "expr => expr => expr => bool" where @@ -160,7 +160,7 @@ | "d x X (Num 1)" | "d (Num n) X (Num 0)" -text {* +text \ ops8(E) :- d((x + num(1)) * ((^(x, 2) + num(2)) * (^(x, 3) + num(3))), x, E). @@ -172,7 +172,7 @@ times10(E) :- d(((((((((x * x) * x) * x) * x) * x) * x) * x) * x) * x, x, E) -*} +\ inductive ops8 :: "expr => bool" where @@ -195,7 +195,7 @@ values_prolog "{e. log10 e}" values_prolog "{e. times10 e}" -section {* Example negation *} +section \Example negation\ datatype abc = A | B | C @@ -203,13 +203,13 @@ where "y \ B \ notB y" -setup {* Code_Prolog.map_code_options (K +setup \Code_Prolog.map_code_options (K {ensure_groundness = true, limit_globally = NONE, limited_types = [], limited_predicates = [], replacing = [], - manual_reorder = []}) *} + manual_reorder = []})\ values_prolog 2 "{y. notB y}" @@ -219,7 +219,7 @@ values_prolog 5 "{y. notAB y}" -section {* Example prolog conform variable names *} +section \Example prolog conform variable names\ inductive equals :: "abc => abc => bool" where