# HG changeset patch # User bulwahn # Date 1256630765 -3600 # Node ID 95186fb5653c9c86f25a278e8e0b5abe9ff3d7ed # Parent b350516cb1f96473f607309b4afbdcfed0d5ce03 added examples for quickcheck prototype diff -r b350516cb1f9 -r 95186fb5653c src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Tue Oct 27 09:03:56 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Tue Oct 27 09:06:05 2009 +0100 @@ -13,11 +13,28 @@ structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK = struct +open Predicate_Compile_Aux; + val test_ref = Unsynchronized.ref (NONE : (unit -> int -> int * int -> term list Predicate.pred * (int * int)) option) val target = "Quickcheck" +val options = Options { + expected_modes = NONE, + show_steps = true, + show_intermediate_results = true, + show_proof_trace = false, + show_modes = true, + show_mode_inference = false, + show_compilation = true, + skip_proof = false, + + inductify = false, + rpred = false, + depth_limited = false +} + fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns) val mk_randompredT = #mk_predT (dest_compfuns Predicate_Compile_Core.randompred_compfuns) @@ -63,18 +80,18 @@ val _ = tracing (Display.string_of_thm ctxt' intro) val thy'' = thy' |> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm intro) - |> Predicate_Compile.preprocess Predicate_Compile_Aux.default_options full_constname - |> Predicate_Compile_Core.add_equations Predicate_Compile_Aux.default_options [full_constname] + |> Predicate_Compile.preprocess options full_constname + |> Predicate_Compile_Core.add_equations options [full_constname] (* |> Predicate_Compile_Core.add_depth_limited_equations Predicate_Compile_Aux.default_options [full_constname]*) - |> Predicate_Compile_Core.add_quickcheck_equations Predicate_Compile_Aux.default_options [full_constname] + |> Predicate_Compile_Core.add_quickcheck_equations options [full_constname] val depth_limited_modes = Predicate_Compile_Core.depth_limited_modes_of thy'' full_constname val modes = Predicate_Compile_Core.generator_modes_of thy'' full_constname val prog = if member (op =) modes ([], []) then let val name = Predicate_Compile_Core.generator_name_of thy'' full_constname ([], []) - val T = [@{typ bool}, @{typ code_numeral}] ---> (mk_randompredT (HOLogic.mk_tupleT (map snd vs'))) - in Const (name, T) $ @{term True} $ Bound 0 end + val T = [@{typ code_numeral}] ---> (mk_randompredT (HOLogic.mk_tupleT (map snd vs'))) + in Const (name, T) $ Bound 0 end (*else if member (op =) depth_limited_modes ([], []) then let val name = Predicate_Compile_Core.depth_limited_function_name_of thy'' full_constname ([], []) diff -r b350516cb1f9 -r 95186fb5653c src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy Tue Oct 27 09:06:05 2009 +0100 @@ -0,0 +1,118 @@ +theory Predicate_Compile_Quickcheck_ex +imports Predicate_Compile_Quickcheck + Predicate_Compile_Alternative_Defs +begin + +section {* Sets *} + + +section {* Context Free Grammar *} + +datatype alphabet = a | b + +inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where + "[] \ S\<^isub>1" +| "w \ A\<^isub>1 \ b # w \ S\<^isub>1" +| "w \ B\<^isub>1 \ a # w \ S\<^isub>1" +| "w \ S\<^isub>1 \ a # w \ A\<^isub>1" +| "w \ S\<^isub>1 \ b # w \ S\<^isub>1" +| "\v \ B\<^isub>1; v \ B\<^isub>1\ \ a # v @ w \ B\<^isub>1" + +theorem S\<^isub>1_sound: +"w \ S\<^isub>1p \ length [x \ w. x = a] = length [x \ w. x = b]" +(*quickcheck[generator=predicate_compile, size=15]*) +oops + + +inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where + "[] \ S\<^isub>2" +| "w \ A\<^isub>2 \ b # w \ S\<^isub>2" +| "w \ B\<^isub>2 \ a # w \ S\<^isub>2" +| "w \ S\<^isub>2 \ a # w \ A\<^isub>2" +| "w \ S\<^isub>2 \ b # w \ B\<^isub>2" +| "\v \ B\<^isub>2; v \ B\<^isub>2\ \ a # v @ w \ B\<^isub>2" + +code_pred [inductify, rpred] S\<^isub>2 . +thm S\<^isub>2.rpred_equation +thm A\<^isub>2.rpred_equation +thm B\<^isub>2.rpred_equation + +values [random] 10 "{x. S\<^isub>2 x}" + +lemma "w \ S\<^isub>2 ==> w \ [] ==> w \ [b, a] ==> w \ {}" +quickcheck[generator=predicate_compile] +oops + +lemma "[x <- w. x = a] = []" +quickcheck[generator=predicate_compile] +oops + + +lemma "length ([x \ w. x = a]) = (0::nat)" +(*quickcheck[generator=predicate_compile]*) +oops + + + +lemma +"w \ S\<^isub>2 ==> length [x \ w. x = a] < Suc (Suc 0)" +(*quickcheck[generator=predicate_compile]*) +oops + + +theorem S\<^isub>2_sound: +"w \ S\<^isub>2 \ length [x \ w. x = a] = length [x \ w. x = b]" +(*quickcheck[generator=predicate_compile, size=15, iterations=100]*) +oops + +inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where + "[] \ S\<^isub>3" +| "w \ A\<^isub>3 \ b # w \ S\<^isub>3" +| "w \ B\<^isub>3 \ a # w \ S\<^isub>3" +| "w \ S\<^isub>3 \ a # w \ A\<^isub>3" +| "w \ S\<^isub>3 \ b # w \ B\<^isub>3" +| "\v \ B\<^isub>3; w \ B\<^isub>3\ \ a # v @ w \ B\<^isub>3" + +code_pred [inductify] S\<^isub>3 . +thm S\<^isub>3.equation + +values 10 "{x. S\<^isub>3 x}" + +lemma S\<^isub>3_sound: +"w \ S\<^isub>3 \ length [x \ w. x = a] = length [x \ w. x = b]" +(*quickcheck[generator=predicate_compile, size=10, iterations=1]*) +oops + + +lemma "\ (length w > 2) \ \ (length [x \ w. x = a] = length [x \ w. x = b])" +(*quickcheck[size=10, generator = pred_compile]*) +oops + +theorem S\<^isub>3_complete: +"length [x \ w. x = a] = length [x \ w. b = x] \ w \ S\<^isub>3" +(*quickcheck[generator=SML]*) +(*quickcheck[generator=predicate_compile, size=10, iterations=100]*) +oops + + +inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where + "[] \ S\<^isub>4" +| "w \ A\<^isub>4 \ b # w \ S\<^isub>4" +| "w \ B\<^isub>4 \ a # w \ S\<^isub>4" +| "w \ S\<^isub>4 \ a # w \ A\<^isub>4" +| "\v \ A\<^isub>4; w \ A\<^isub>4\ \ b # v @ w \ A\<^isub>4" +| "w \ S\<^isub>4 \ b # w \ B\<^isub>4" +| "\v \ B\<^isub>4; w \ B\<^isub>4\ \ a # v @ w \ B\<^isub>4" + +theorem S\<^isub>4_sound: +"w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" +(*quickcheck[generator = predicate_compile, size=2, iterations=1]*) +oops + +theorem S\<^isub>4_complete: +"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^isub>4" +(*quickcheck[generator = pred_compile, size=5, iterations=1]*) +oops + + +end \ No newline at end of file