# HG changeset patch # User wenzelm # Date 1361791070 -3600 # Node ID 9c8d63b4b6be2216713c78fe0c607fdafa88c75c # Parent e8d2ecf6aaac7f357c6c7597da10d95ab2cd1f0e prefer stateless 'ML_val' for tests; diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Decision_Procs/Cooper.thy --- a/src/HOL/Decision_Procs/Cooper.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Decision_Procs/Cooper.thy Mon Feb 25 12:17:50 2013 +0100 @@ -1994,7 +1994,7 @@ pa (E (A (Imp (Ge (Sub (Bound 0) (Bound 1))) (E (E (Eq (Sub (Add (Mul 3 (Bound 1)) (Mul 5 (Bound 0))) (Bound 2))))))))" -ML {* @{code cooper_test} () *} +ML_val {* @{code cooper_test} () *} (*code_reflect Cooper_Procedure functions pa diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Decision_Procs/Ferrack.thy --- a/src/HOL/Decision_Procs/Ferrack.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Decision_Procs/Ferrack.thy Mon Feb 25 12:17:50 2013 +0100 @@ -1917,7 +1917,7 @@ "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0))) (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))" -ML {* @{code ferrack_test} () *} +ML_val {* @{code ferrack_test} () *} oracle linr_oracle = {* let diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Mon Feb 25 12:17:50 2013 +0100 @@ -5502,17 +5502,15 @@ definition "problem4 = E (And (Ge (Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg (Bound 0))))))" -ML {* - Par_List.map (fn e => e ()) - [fn () => @{code mircfrqe} @{code problem1}, - fn () => @{code mirlfrqe} @{code problem1}, - fn () => @{code mircfrqe} @{code problem2}, - fn () => @{code mirlfrqe} @{code problem2}, - fn () => @{code mircfrqe} @{code problem3}, - fn () => @{code mirlfrqe} @{code problem3}, - fn () => @{code mircfrqe} @{code problem4}, - fn () => @{code mirlfrqe} @{code problem4}] -*} +ML_val {* @{code mircfrqe} @{code problem1} *} +ML_val {* @{code mirlfrqe} @{code problem1} *} +ML_val {* @{code mircfrqe} @{code problem2} *} +ML_val {* @{code mirlfrqe} @{code problem2} *} +ML_val {* @{code mircfrqe} @{code problem3} *} +ML_val {* @{code mirlfrqe} @{code problem3} *} +ML_val {* @{code mircfrqe} @{code problem4} *} +ML_val {* @{code mirlfrqe} @{code problem4} *} + (*code_reflect Mir functions mircfrqe mirlfrqe diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy --- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Mon Feb 25 12:17:50 2013 +0100 @@ -662,7 +662,7 @@ qsort a }" -ML {* @{code example} () *} +ML_val {* @{code example} () *} export_code qsort checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala Scala_imp diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Imperative_HOL/ex/Linked_Lists.thy --- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Feb 25 12:17:50 2013 +0100 @@ -996,9 +996,9 @@ code_reserved SML upto -ML {* @{code test_1} () *} -ML {* @{code test_2} () *} -ML {* @{code test_3} () *} +ML_val {* @{code test_1} () *} +ML_val {* @{code test_2} () *} +ML_val {* @{code test_3} () *} export_code test_1 test_2 test_3 checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala Scala_imp diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/List.thy --- a/src/HOL/List.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/List.thy Mon Feb 25 12:17:50 2013 +0100 @@ -444,7 +444,7 @@ in [(@{syntax_const "_listcompr"}, lc_tr)] end *} -ML {* +ML_val {* let val read = Syntax.read_term @{context}; fun check s1 s2 = read s1 aconv read s2 orelse error ("Check failed: " ^ quote s1); diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/MicroJava/BV/BVExample.thy Mon Feb 25 12:17:50 2013 +0100 @@ -480,7 +480,7 @@ definition test2 where "test2 = test_kil E test_name [] (PrimT Void) 3 2 [] make_list_ins" -ML {* +ML_val {* @{code test1}; @{code test2}; *} diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/MicroJava/J/JListExample.thy --- a/src/HOL/MicroJava/J/JListExample.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/MicroJava/J/JListExample.thy Mon Feb 25 12:17:50 2013 +0100 @@ -151,7 +151,7 @@ "test = exec_i_i_i_o example_prg (Norm (empty, empty)) E" by(auto intro: exec_i_i_i_oI intro!: pred_eqI elim: exec_i_i_i_oE simp add: test_def) -ML {* +ML_val {* val SOME ((_, (heap, locs)), _) = Predicate.yield @{code test}; locs @{code l1_name}; locs @{code l2_name}; diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/MicroJava/JVM/JVMListExample.thy --- a/src/HOL/MicroJava/JVM/JVMListExample.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/MicroJava/JVM/JVMListExample.thy Mon Feb 25 12:17:50 2013 +0100 @@ -136,7 +136,7 @@ definition "test = exec (E, start_state E test_name makelist_name)" -ML {* +ML_val {* @{code test}; @{code exec} (@{code E}, @{code the} it); @{code exec} (@{code E}, @{code the} it); diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Nitpick_Examples/Mini_Nits.thy --- a/src/HOL/Nitpick_Examples/Mini_Nits.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Nitpick_Examples/Mini_Nits.thy Mon Feb 25 12:17:50 2013 +0100 @@ -25,88 +25,88 @@ val unknown = expect "unknown" *} -ML {* genuine 1 @{prop "x = Not"} *} -ML {* none 1 @{prop "\x. x = Not"} *} -ML {* none 1 @{prop "\ False"} *} -ML {* genuine 1 @{prop "\ True"} *} -ML {* none 1 @{prop "\ \ b \ b"} *} -ML {* none 1 @{prop True} *} -ML {* genuine 1 @{prop False} *} -ML {* genuine 1 @{prop "True \ False"} *} -ML {* none 1 @{prop "True \ \ False"} *} -ML {* none 5 @{prop "\x. x = x"} *} -ML {* none 5 @{prop "\x. x = x"} *} -ML {* none 1 @{prop "\x. x = y"} *} -ML {* genuine 2 @{prop "\x. x = y"} *} -ML {* none 2 @{prop "\x. x = y"} *} -ML {* none 2 @{prop "\x\'a \ 'a. x = x"} *} -ML {* none 2 @{prop "\x\'a \ 'a. x = y"} *} -ML {* genuine 2 @{prop "\x\'a \ 'a. x = y"} *} -ML {* none 2 @{prop "\x\'a \ 'a. x = y"} *} -ML {* none 1 @{prop "All = Ex"} *} -ML {* genuine 2 @{prop "All = Ex"} *} -ML {* none 1 @{prop "All P = Ex P"} *} -ML {* genuine 2 @{prop "All P = Ex P"} *} -ML {* none 5 @{prop "x = y \ P x = P y"} *} -ML {* none 5 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} -ML {* none 2 @{prop "(x\'a \ 'a) = y \ P x y = P y x"} *} -ML {* none 5 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} -ML {* none 2 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} -ML {* none 2 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} -ML {* genuine 1 @{prop "(op =) X = Ex"} *} -ML {* none 2 @{prop "\x::'a \ 'a. x = x"} *} -ML {* none 1 @{prop "x = y"} *} -ML {* genuine 1 @{prop "x \ y"} *} -ML {* genuine 2 @{prop "x = y"} *} -ML {* genuine 1 @{prop "X \ Y"} *} -ML {* none 1 @{prop "P \ Q \ Q \ P"} *} -ML {* none 1 @{prop "P \ Q \ P"} *} -ML {* none 1 @{prop "P \ Q \ Q \ P"} *} -ML {* genuine 1 @{prop "P \ Q \ P"} *} -ML {* none 1 @{prop "(P \ Q) \ (\ P \ Q)"} *} -ML {* none 5 @{prop "{a} = {a, a}"} *} -ML {* genuine 2 @{prop "{a} = {a, b}"} *} -ML {* genuine 1 @{prop "{a} \ {a, b}"} *} -ML {* none 5 @{prop "{}\<^sup>+ = {}"} *} -ML {* none 5 @{prop "UNIV\<^sup>+ = UNIV"} *} -ML {* none 5 @{prop "(UNIV \ ('a \ 'b) set) - {} = UNIV"} *} -ML {* none 5 @{prop "{} - (UNIV \ ('a \ 'b) set) = {}"} *} -ML {* none 1 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} -ML {* genuine 2 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} -ML {* none 5 @{prop "a \ c \ {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} -ML {* none 5 @{prop "A \ B = {x. x \ A \ x \ B}"} *} -ML {* none 5 @{prop "A \ B = {x. x \ A \ x \ B}"} *} -ML {* none 5 @{prop "A - B = (\x. A x \ \ B x)"} *} -ML {* none 5 @{prop "\a b. (a, b) = (b, a)"} *} -ML {* genuine 2 @{prop "(a, b) = (b, a)"} *} -ML {* genuine 2 @{prop "(a, b) \ (b, a)"} *} -ML {* none 5 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} -ML {* genuine 2 @{prop "(a\'a \ 'a, b) = (b, a)"} *} -ML {* none 5 @{prop "\a b\'a \ 'a \ 'a. (a, b) = (b, a)"} *} -ML {* genuine 2 @{prop "(a\'a \ 'a \ 'a, b) \ (b, a)"} *} -ML {* none 5 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} -ML {* genuine 1 @{prop "(a\'a \ 'a, b) \ (b, a)"} *} -ML {* none 5 @{prop "fst (a, b) = a"} *} -ML {* none 1 @{prop "fst (a, b) = b"} *} -ML {* genuine 2 @{prop "fst (a, b) = b"} *} -ML {* genuine 2 @{prop "fst (a, b) \ b"} *} -ML {* genuine 2 @{prop "f ((x, z), y) = (x, z)"} *} -ML {* none 2 @{prop "(ALL x. f x = fst x) \ f ((x, z), y) = (x, z)"} *} -ML {* none 5 @{prop "snd (a, b) = b"} *} -ML {* none 1 @{prop "snd (a, b) = a"} *} -ML {* genuine 2 @{prop "snd (a, b) = a"} *} -ML {* genuine 2 @{prop "snd (a, b) \ a"} *} -ML {* genuine 1 @{prop P} *} -ML {* genuine 1 @{prop "(\x. P) a"} *} -ML {* genuine 1 @{prop "(\x y z. P y x z) a b c"} *} -ML {* none 5 @{prop "\f. f = (\x. x) \ f y = y"} *} -ML {* genuine 1 @{prop "\f. f p \ p \ (\a b. f (a, b) = (a, b))"} *} -ML {* none 2 @{prop "\f. \a b. f (a, b) = (a, b)"} *} -ML {* none 3 @{prop "f = (\a b. (b, a)) \ f x y = (y, x)"} *} -ML {* genuine 2 @{prop "f = (\a b. (b, a)) \ f x y = (x, y)"} *} -ML {* none 5 @{prop "f = (\x. f x)"} *} -ML {* none 5 @{prop "f = (\x. f x\'a \ bool)"} *} -ML {* none 5 @{prop "f = (\x y. f x y)"} *} -ML {* none 5 @{prop "f = (\x y. f x y\'a \ bool)"} *} +ML_val {* genuine 1 @{prop "x = Not"} *} +ML_val {* none 1 @{prop "\x. x = Not"} *} +ML_val {* none 1 @{prop "\ False"} *} +ML_val {* genuine 1 @{prop "\ True"} *} +ML_val {* none 1 @{prop "\ \ b \ b"} *} +ML_val {* none 1 @{prop True} *} +ML_val {* genuine 1 @{prop False} *} +ML_val {* genuine 1 @{prop "True \ False"} *} +ML_val {* none 1 @{prop "True \ \ False"} *} +ML_val {* none 5 @{prop "\x. x = x"} *} +ML_val {* none 5 @{prop "\x. x = x"} *} +ML_val {* none 1 @{prop "\x. x = y"} *} +ML_val {* genuine 2 @{prop "\x. x = y"} *} +ML_val {* none 2 @{prop "\x. x = y"} *} +ML_val {* none 2 @{prop "\x\'a \ 'a. x = x"} *} +ML_val {* none 2 @{prop "\x\'a \ 'a. x = y"} *} +ML_val {* genuine 2 @{prop "\x\'a \ 'a. x = y"} *} +ML_val {* none 2 @{prop "\x\'a \ 'a. x = y"} *} +ML_val {* none 1 @{prop "All = Ex"} *} +ML_val {* genuine 2 @{prop "All = Ex"} *} +ML_val {* none 1 @{prop "All P = Ex P"} *} +ML_val {* genuine 2 @{prop "All P = Ex P"} *} +ML_val {* none 5 @{prop "x = y \ P x = P y"} *} +ML_val {* none 5 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} +ML_val {* none 2 @{prop "(x\'a \ 'a) = y \ P x y = P y x"} *} +ML_val {* none 5 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} +ML_val {* none 2 @{prop "(x\'a \ 'a) = y \ P x = P y"} *} +ML_val {* none 2 @{prop "\x\'a \ 'a. x = y \ P x = P y"} *} +ML_val {* genuine 1 @{prop "(op =) X = Ex"} *} +ML_val {* none 2 @{prop "\x::'a \ 'a. x = x"} *} +ML_val {* none 1 @{prop "x = y"} *} +ML_val {* genuine 1 @{prop "x \ y"} *} +ML_val {* genuine 2 @{prop "x = y"} *} +ML_val {* genuine 1 @{prop "X \ Y"} *} +ML_val {* none 1 @{prop "P \ Q \ Q \ P"} *} +ML_val {* none 1 @{prop "P \ Q \ P"} *} +ML_val {* none 1 @{prop "P \ Q \ Q \ P"} *} +ML_val {* genuine 1 @{prop "P \ Q \ P"} *} +ML_val {* none 1 @{prop "(P \ Q) \ (\ P \ Q)"} *} +ML_val {* none 5 @{prop "{a} = {a, a}"} *} +ML_val {* genuine 2 @{prop "{a} = {a, b}"} *} +ML_val {* genuine 1 @{prop "{a} \ {a, b}"} *} +ML_val {* none 5 @{prop "{}\<^sup>+ = {}"} *} +ML_val {* none 5 @{prop "UNIV\<^sup>+ = UNIV"} *} +ML_val {* none 5 @{prop "(UNIV \ ('a \ 'b) set) - {} = UNIV"} *} +ML_val {* none 5 @{prop "{} - (UNIV \ ('a \ 'b) set) = {}"} *} +ML_val {* none 1 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} +ML_val {* genuine 2 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} +ML_val {* none 5 @{prop "a \ c \ {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *} +ML_val {* none 5 @{prop "A \ B = {x. x \ A \ x \ B}"} *} +ML_val {* none 5 @{prop "A \ B = {x. x \ A \ x \ B}"} *} +ML_val {* none 5 @{prop "A - B = (\x. A x \ \ B x)"} *} +ML_val {* none 5 @{prop "\a b. (a, b) = (b, a)"} *} +ML_val {* genuine 2 @{prop "(a, b) = (b, a)"} *} +ML_val {* genuine 2 @{prop "(a, b) \ (b, a)"} *} +ML_val {* none 5 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} +ML_val {* genuine 2 @{prop "(a\'a \ 'a, b) = (b, a)"} *} +ML_val {* none 5 @{prop "\a b\'a \ 'a \ 'a. (a, b) = (b, a)"} *} +ML_val {* genuine 2 @{prop "(a\'a \ 'a \ 'a, b) \ (b, a)"} *} +ML_val {* none 5 @{prop "\a b\'a \ 'a. (a, b) = (b, a)"} *} +ML_val {* genuine 1 @{prop "(a\'a \ 'a, b) \ (b, a)"} *} +ML_val {* none 5 @{prop "fst (a, b) = a"} *} +ML_val {* none 1 @{prop "fst (a, b) = b"} *} +ML_val {* genuine 2 @{prop "fst (a, b) = b"} *} +ML_val {* genuine 2 @{prop "fst (a, b) \ b"} *} +ML_val {* genuine 2 @{prop "f ((x, z), y) = (x, z)"} *} +ML_val {* none 2 @{prop "(ALL x. f x = fst x) \ f ((x, z), y) = (x, z)"} *} +ML_val {* none 5 @{prop "snd (a, b) = b"} *} +ML_val {* none 1 @{prop "snd (a, b) = a"} *} +ML_val {* genuine 2 @{prop "snd (a, b) = a"} *} +ML_val {* genuine 2 @{prop "snd (a, b) \ a"} *} +ML_val {* genuine 1 @{prop P} *} +ML_val {* genuine 1 @{prop "(\x. P) a"} *} +ML_val {* genuine 1 @{prop "(\x y z. P y x z) a b c"} *} +ML_val {* none 5 @{prop "\f. f = (\x. x) \ f y = y"} *} +ML_val {* genuine 1 @{prop "\f. f p \ p \ (\a b. f (a, b) = (a, b))"} *} +ML_val {* none 2 @{prop "\f. \a b. f (a, b) = (a, b)"} *} +ML_val {* none 3 @{prop "f = (\a b. (b, a)) \ f x y = (y, x)"} *} +ML_val {* genuine 2 @{prop "f = (\a b. (b, a)) \ f x y = (x, y)"} *} +ML_val {* none 5 @{prop "f = (\x. f x)"} *} +ML_val {* none 5 @{prop "f = (\x. f x\'a \ bool)"} *} +ML_val {* none 5 @{prop "f = (\x y. f x y)"} *} +ML_val {* none 5 @{prop "f = (\x y. f x y\'a \ bool)"} *} end diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Nitpick_Examples/Mono_Nits.thy --- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Mon Feb 25 12:17:50 2013 +0100 @@ -67,78 +67,78 @@ ML {* Nitpick_Mono.trace := false *} -ML {* const @{term "A\('a\'b)"} *} -ML {* const @{term "(A\'a set) = A"} *} -ML {* const @{term "(A\'a set set) = A"} *} -ML {* const @{term "(\x\'a set. a \ x)"} *} -ML {* const @{term "{{a\'a}} = C"} *} -ML {* const @{term "{f\'a\nat} = {g\'a\nat}"} *} -ML {* const @{term "A \ (B\'a set)"} *} -ML {* const @{term "\A B x\'a. A x \ B x"} *} -ML {* const @{term "P (a\'a)"} *} -ML {* const @{term "\a\'a. b (c (d\'a)) (e\'a) (f\'a)"} *} -ML {* const @{term "\A\'a set. a \ A"} *} -ML {* const @{term "\A\'a set. P A"} *} -ML {* const @{term "P \ Q"} *} -ML {* const @{term "A \ B = (C\'a set)"} *} -ML {* const @{term "(\A B x\'a. A x \ B x) A B = C"} *} -ML {* const @{term "(if P then (A\'a set) else B) = C"} *} -ML {* const @{term "let A = (C\'a set) in A \ B"} *} -ML {* const @{term "THE x\'b. P x"} *} -ML {* const @{term "(\x\'a. False)"} *} -ML {* const @{term "(\x\'a. True)"} *} -ML {* const @{term "(\x\'a. False) = (\x\'a. False)"} *} -ML {* const @{term "(\x\'a. True) = (\x\'a. True)"} *} -ML {* const @{term "Let (a\'a) A"} *} -ML {* const @{term "A (a\'a)"} *} -ML {* const @{term "insert (a\'a) A = B"} *} -ML {* const @{term "- (A\'a set)"} *} -ML {* const @{term "finite (A\'a set)"} *} -ML {* const @{term "\ finite (A\'a set)"} *} -ML {* const @{term "finite (A\'a set set)"} *} -ML {* const @{term "\a\'a. A a \ \ B a"} *} -ML {* const @{term "A < (B\'a set)"} *} -ML {* const @{term "A \ (B\'a set)"} *} -ML {* const @{term "[a\'a]"} *} -ML {* const @{term "[a\'a set]"} *} -ML {* const @{term "[A \ (B\'a set)]"} *} -ML {* const @{term "[A \ (B\'a set)] = [C]"} *} -ML {* const @{term "{(\x\'a. x = a)} = C"} *} -ML {* const @{term "(\a\'a. \ A a) = B"} *} -ML {* const @{prop "\F f g (h\'a set). F f \ F g \ \ f a \ g a \ \ f a"} *} -ML {* const @{term "\A B x\'a. A x \ B x \ A = B"} *} -ML {* const @{term "p = (\(x\'a) (y\'a). P x \ \ Q y)"} *} -ML {* const @{term "p = (\(x\'a) (y\'a). p x y \ bool)"} *} -ML {* const @{term "p = (\A B x. A x \ \ B x) (\x. True) (\y. x \ y)"} *} -ML {* const @{term "p = (\y. x \ y)"} *} -ML {* const @{term "(\x. (p\'a\bool\bool) x False)"} *} -ML {* const @{term "(\x y. (p\'a\'a\bool\bool) x y False)"} *} -ML {* const @{term "f = (\x\'a. P x \ Q x)"} *} -ML {* const @{term "\a\'a. P a"} *} +ML_val {* const @{term "A\('a\'b)"} *} +ML_val {* const @{term "(A\'a set) = A"} *} +ML_val {* const @{term "(A\'a set set) = A"} *} +ML_val {* const @{term "(\x\'a set. a \ x)"} *} +ML_val {* const @{term "{{a\'a}} = C"} *} +ML_val {* const @{term "{f\'a\nat} = {g\'a\nat}"} *} +ML_val {* const @{term "A \ (B\'a set)"} *} +ML_val {* const @{term "\A B x\'a. A x \ B x"} *} +ML_val {* const @{term "P (a\'a)"} *} +ML_val {* const @{term "\a\'a. b (c (d\'a)) (e\'a) (f\'a)"} *} +ML_val {* const @{term "\A\'a set. a \ A"} *} +ML_val {* const @{term "\A\'a set. P A"} *} +ML_val {* const @{term "P \ Q"} *} +ML_val {* const @{term "A \ B = (C\'a set)"} *} +ML_val {* const @{term "(\A B x\'a. A x \ B x) A B = C"} *} +ML_val {* const @{term "(if P then (A\'a set) else B) = C"} *} +ML_val {* const @{term "let A = (C\'a set) in A \ B"} *} +ML_val {* const @{term "THE x\'b. P x"} *} +ML_val {* const @{term "(\x\'a. False)"} *} +ML_val {* const @{term "(\x\'a. True)"} *} +ML_val {* const @{term "(\x\'a. False) = (\x\'a. False)"} *} +ML_val {* const @{term "(\x\'a. True) = (\x\'a. True)"} *} +ML_val {* const @{term "Let (a\'a) A"} *} +ML_val {* const @{term "A (a\'a)"} *} +ML_val {* const @{term "insert (a\'a) A = B"} *} +ML_val {* const @{term "- (A\'a set)"} *} +ML_val {* const @{term "finite (A\'a set)"} *} +ML_val {* const @{term "\ finite (A\'a set)"} *} +ML_val {* const @{term "finite (A\'a set set)"} *} +ML_val {* const @{term "\a\'a. A a \ \ B a"} *} +ML_val {* const @{term "A < (B\'a set)"} *} +ML_val {* const @{term "A \ (B\'a set)"} *} +ML_val {* const @{term "[a\'a]"} *} +ML_val {* const @{term "[a\'a set]"} *} +ML_val {* const @{term "[A \ (B\'a set)]"} *} +ML_val {* const @{term "[A \ (B\'a set)] = [C]"} *} +ML_val {* const @{term "{(\x\'a. x = a)} = C"} *} +ML_val {* const @{term "(\a\'a. \ A a) = B"} *} +ML_val {* const @{prop "\F f g (h\'a set). F f \ F g \ \ f a \ g a \ \ f a"} *} +ML_val {* const @{term "\A B x\'a. A x \ B x \ A = B"} *} +ML_val {* const @{term "p = (\(x\'a) (y\'a). P x \ \ Q y)"} *} +ML_val {* const @{term "p = (\(x\'a) (y\'a). p x y \ bool)"} *} +ML_val {* const @{term "p = (\A B x. A x \ \ B x) (\x. True) (\y. x \ y)"} *} +ML_val {* const @{term "p = (\y. x \ y)"} *} +ML_val {* const @{term "(\x. (p\'a\bool\bool) x False)"} *} +ML_val {* const @{term "(\x y. (p\'a\'a\bool\bool) x y False)"} *} +ML_val {* const @{term "f = (\x\'a. P x \ Q x)"} *} +ML_val {* const @{term "\a\'a. P a"} *} -ML {* nonconst @{term "\P (a\'a). P a"} *} -ML {* nonconst @{term "THE x\'a. P x"} *} -ML {* nonconst @{term "SOME x\'a. P x"} *} -ML {* nonconst @{term "(\A B x\'a. A x \ B x) = myunion"} *} -ML {* nonconst @{term "(\x\'a. False) = (\x\'a. True)"} *} -ML {* nonconst @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} +ML_val {* nonconst @{term "\P (a\'a). P a"} *} +ML_val {* nonconst @{term "THE x\'a. P x"} *} +ML_val {* nonconst @{term "SOME x\'a. P x"} *} +ML_val {* nonconst @{term "(\A B x\'a. A x \ B x) = myunion"} *} +ML_val {* nonconst @{term "(\x\'a. False) = (\x\'a. True)"} *} +ML_val {* nonconst @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} -ML {* mono @{prop "Q (\x\'a set. P x)"} *} -ML {* mono @{prop "P (a\'a)"} *} -ML {* mono @{prop "{a} = {b\'a}"} *} -ML {* mono @{prop "(\x. x = a) = (\y. y = (b\'a))"} *} -ML {* mono @{prop "(a\'a) \ P \ P \ P = P"} *} -ML {* mono @{prop "\F\'a set set. P"} *} -ML {* mono @{prop "\ (\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h)"} *} -ML {* mono @{prop "\ Q (\x\'a set. P x)"} *} -ML {* mono @{prop "\ (\x\'a. P x)"} *} -ML {* mono @{prop "myall P = (P = (\x\'a. True))"} *} -ML {* mono @{prop "myall P = (P = (\x\'a. False))"} *} -ML {* mono @{prop "\x\'a. P x"} *} -ML {* mono @{term "(\A B x\'a. A x \ B x) \ myunion"} *} +ML_val {* mono @{prop "Q (\x\'a set. P x)"} *} +ML_val {* mono @{prop "P (a\'a)"} *} +ML_val {* mono @{prop "{a} = {b\'a}"} *} +ML_val {* mono @{prop "(\x. x = a) = (\y. y = (b\'a))"} *} +ML_val {* mono @{prop "(a\'a) \ P \ P \ P = P"} *} +ML_val {* mono @{prop "\F\'a set set. P"} *} +ML_val {* mono @{prop "\ (\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h)"} *} +ML_val {* mono @{prop "\ Q (\x\'a set. P x)"} *} +ML_val {* mono @{prop "\ (\x\'a. P x)"} *} +ML_val {* mono @{prop "myall P = (P = (\x\'a. True))"} *} +ML_val {* mono @{prop "myall P = (P = (\x\'a. False))"} *} +ML_val {* mono @{prop "\x\'a. P x"} *} +ML_val {* mono @{term "(\A B x\'a. A x \ B x) \ myunion"} *} -ML {* nonmono @{prop "A = (\x::'a. True) \ A = (\x. False)"} *} -ML {* nonmono @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} +ML_val {* nonmono @{prop "A = (\x::'a. True) \ A = (\x. False)"} *} +ML_val {* nonmono @{prop "\F f g (h\'a set). F f \ F g \ \ a \ f \ a \ g \ F h"} *} ML {* val preproc_timeout = SOME (seconds 5.0) @@ -200,12 +200,12 @@ *} (* -ML {* check_theory @{theory AVL2} *} -ML {* check_theory @{theory Fun} *} -ML {* check_theory @{theory Huffman} *} -ML {* check_theory @{theory List} *} -ML {* check_theory @{theory Map} *} -ML {* check_theory @{theory Relation} *} +ML_val {* check_theory @{theory AVL2} *} +ML_val {* check_theory @{theory Fun} *} +ML_val {* check_theory @{theory Huffman} *} +ML_val {* check_theory @{theory List} *} +ML_val {* check_theory @{theory Map} *} +ML_val {* check_theory @{theory Relation} *} *) ML {* getenv "ISABELLE_TMP" *} diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy --- a/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Mon Feb 25 12:17:50 2013 +0100 @@ -15,7 +15,7 @@ "greater_than_index xs = (\i x. nth_el' xs i = Some x --> x > i)" code_pred (expected_modes: i => bool) [inductify, skip_proof, specialise] greater_than_index . -ML {* Core_Data.intros_of @{context} @{const_name specialised_nth_el'P} *} +ML_val {* Core_Data.intros_of @{context} @{const_name specialised_nth_el'P} *} thm greater_than_index.equation @@ -44,7 +44,7 @@ thm max_of_my_SucP.equation -ML {* Core_Data.intros_of @{context} @{const_name specialised_max_natP} *} +ML_val {* Core_Data.intros_of @{context} @{const_name specialised_max_natP} *} values "{x. max_of_my_SucP x 6}" diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Proofs/Extraction/Higman_Extraction.thy --- a/src/HOL/Proofs/Extraction/Higman_Extraction.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Proofs/Extraction/Higman_Extraction.thy Mon Feb 25 12:17:50 2013 +0100 @@ -88,7 +88,7 @@ | "f2 (Suc (Suc 0)) = [B, A]" | "f2 _ = []" -ML {* +ML_val {* local val higman_idx = @{code higman_idx}; val g1 = @{code g1}; diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Proofs/Extraction/Pigeonhole.thy --- a/src/HOL/Proofs/Extraction/Pigeonhole.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Proofs/Extraction/Pigeonhole.thy Mon Feb 25 12:17:50 2013 +0100 @@ -243,14 +243,14 @@ definition "test'' u = pigeonhole 8 (List.nth [0, 1, 2, 3, 4, 5, 6, 3, 7, 8])" -ML "timeit (@{code test} 10)" -ML "timeit (@{code test'} 10)" -ML "timeit (@{code test} 20)" -ML "timeit (@{code test'} 20)" -ML "timeit (@{code test} 25)" -ML "timeit (@{code test'} 25)" -ML "timeit (@{code test} 500)" -ML "timeit @{code test''}" +ML_val "timeit (@{code test} 10)" +ML_val "timeit (@{code test'} 10)" +ML_val "timeit (@{code test} 20)" +ML_val "timeit (@{code test'} 20)" +ML_val "timeit (@{code test} 25)" +ML_val "timeit (@{code test'} 25)" +ML_val "timeit (@{code test} 500)" +ML_val "timeit @{code test''}" end diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Proofs/Extraction/Warshall.thy --- a/src/HOL/Proofs/Extraction/Warshall.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Proofs/Extraction/Warshall.thy Mon Feb 25 12:17:50 2013 +0100 @@ -256,6 +256,6 @@ @{thm [display] warshall_correctness [no_vars]} *} -ML "@{code warshall}" +ML_val "@{code warshall}" end diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Quickcheck_Examples/Hotel_Example.thy --- a/src/HOL/Quickcheck_Examples/Hotel_Example.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Quickcheck_Examples/Hotel_Example.thy Mon Feb 25 12:17:50 2013 +0100 @@ -92,7 +92,7 @@ setup {* Predicate_Compile_Data.ignore_consts [@{const_name Set.member}, @{const_name "issued"}, @{const_name "cards"}, @{const_name "isin"}, @{const_name Collect}, @{const_name insert}] *} -ML {* Core_Data.force_modes_and_compilations *} +ML_val {* Core_Data.force_modes_and_compilations *} fun find_first :: "('a => 'b option) => 'a list => 'b option" where diff -r e8d2ecf6aaac -r 9c8d63b4b6be src/HOL/Quickcheck_Examples/Quickcheck_Interfaces.thy --- a/src/HOL/Quickcheck_Examples/Quickcheck_Interfaces.thy Mon Feb 25 12:17:11 2013 +0100 +++ b/src/HOL/Quickcheck_Examples/Quickcheck_Interfaces.thy Mon Feb 25 12:17:50 2013 +0100 @@ -33,16 +33,16 @@ fun check_upto f i j = if i > j then true else f i andalso check_upto f (i + 1) j *} -ML {* +ML_val {* map (fn test => check_upto test 0 1) testers *} -ML {* +ML_val {* map (fn test => check_upto test 0 2) testers *} -ML {* +ML_val {* map (fn test => check_upto test 0 3) testers *} -ML {* +ML_val {* map (fn test => check_upto test 0 7) testers *}