src/HOL/ex/Predicate_Compile_Quickcheck.thy
author wenzelm
Sat, 13 Mar 2010 15:12:47 +0100
changeset 35745 1416f568b2b6
parent 35537 59dd6be5834c
child 35951 5ad0af66b3c6
permissions -rw-r--r--
command 'typedef' now works within a local theory context;

(* Author: Lukas Bulwahn, TU Muenchen *)

header {* A Prototype of Quickcheck based on the Predicate Compiler *}

theory Predicate_Compile_Quickcheck
imports "../Predicate_Compile" Quickcheck Predicate_Compile_Alternative_Defs
uses "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
begin

setup {* Quickcheck.add_generator ("predicate_compile_wo_ff", Predicate_Compile_Quickcheck.quickcheck_compile_term false true 8) *}
setup {* Quickcheck.add_generator ("predicate_compile_ff_fs", Predicate_Compile_Quickcheck.quickcheck_compile_term true true 8) *}
setup {* Quickcheck.add_generator ("predicate_compile_ff_nofs", Predicate_Compile_Quickcheck.quickcheck_compile_term true false 8) *}

(*
datatype alphabet = a | b

inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
  "[] \<in> S\<^isub>1"
| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"

ML {* set Toplevel.debug *}

declare mem_def[code_pred_inline] Collect_def[code_pred_inline]

lemma
  "w \<in> S\<^isub>1p \<Longrightarrow> w = []"
quickcheck[generator = predicate_compile, iterations=1]
oops

theorem S\<^isub>1_sound:
"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[generator=predicate_compile, size=15]
oops
*)
end