bulwahn@41922: (* Author: Florian Haftmann & Lukas Bulwahn, TU Muenchen *) haftmann@26265: bulwahn@41922: header {* A simple counterexample generator performing random testing *} haftmann@26265: haftmann@51126: theory Quickcheck_Random bulwahn@40650: imports Random Code_Evaluation Enum haftmann@26265: begin haftmann@26265: haftmann@37751: notation fcomp (infixl "\>" 60) haftmann@37751: notation scomp (infixl "\\" 60) haftmann@31179: bulwahn@45718: setup {* Code_Target.extend_target ("Quickcheck", (Code_Runtime.target, K I)) *} bulwahn@45718: bulwahn@45718: subsection {* Catching Match exceptions *} bulwahn@45718: bulwahn@45801: axiomatization catch_match :: "'a => 'a => 'a" bulwahn@45718: haftmann@52435: code_printing haftmann@52435: constant catch_match \ (Quickcheck) "((_) handle Match => _)" haftmann@31179: haftmann@26265: subsection {* The @{text random} class *} haftmann@26265: haftmann@28335: class random = typerep + haftmann@51143: fixes random :: "natural \ Random.seed \ ('a \ (unit \ term)) \ Random.seed" haftmann@26265: haftmann@26267: haftmann@31260: subsection {* Fundamental and numeric types*} haftmann@31179: haftmann@31179: instantiation bool :: random haftmann@31179: begin haftmann@31179: haftmann@31179: definition haftmann@37751: "random i = Random.range 2 \\ haftmann@32657: (\k. Pair (if k = 0 then Code_Evaluation.valtermify False else Code_Evaluation.valtermify True))" haftmann@31179: haftmann@31179: instance .. haftmann@31179: haftmann@31179: end haftmann@31179: haftmann@31179: instantiation itself :: (typerep) random haftmann@31179: begin haftmann@31179: wenzelm@46975: definition haftmann@51143: random_itself :: "natural \ Random.seed \ ('a itself \ (unit \ term)) \ Random.seed" wenzelm@46975: where "random_itself _ = Pair (Code_Evaluation.valtermify TYPE('a))" haftmann@31179: haftmann@31179: instance .. haftmann@31179: haftmann@31179: end haftmann@31179: haftmann@31483: instantiation char :: random haftmann@31483: begin haftmann@31483: haftmann@31483: definition haftmann@49972: "random _ = Random.select (Enum.enum :: char list) \\ (\c. Pair (c, \u. Code_Evaluation.term_of c))" haftmann@31483: haftmann@31483: instance .. haftmann@31483: haftmann@31483: end haftmann@31483: haftmann@31483: instantiation String.literal :: random haftmann@31483: begin haftmann@31483: haftmann@31483: definition haftmann@32657: "random _ = Pair (STR '''', \u. Code_Evaluation.term_of (STR ''''))" haftmann@31483: haftmann@31483: instance .. haftmann@31483: haftmann@31483: end haftmann@31483: haftmann@31179: instantiation nat :: random haftmann@31179: begin haftmann@31179: haftmann@51143: definition random_nat :: "natural \ Random.seed wenzelm@46975: \ (nat \ (unit \ Code_Evaluation.term)) \ Random.seed" wenzelm@46975: where haftmann@37751: "random_nat i = Random.range (i + 1) \\ (\k. Pair ( haftmann@51143: let n = nat_of_natural k haftmann@32657: in (n, \_. Code_Evaluation.term_of n)))" haftmann@31194: haftmann@31194: instance .. haftmann@31194: haftmann@31194: end haftmann@31194: haftmann@31194: instantiation int :: random haftmann@31194: begin haftmann@31194: haftmann@31194: definition haftmann@37751: "random i = Random.range (2 * i + 1) \\ (\k. Pair ( haftmann@51143: let j = (if k \ i then int (nat_of_natural (k - i)) else - (int (nat_of_natural (i - k)))) haftmann@32657: in (j, \_. Code_Evaluation.term_of j)))" haftmann@31179: haftmann@31179: instance .. haftmann@31179: haftmann@30945: end haftmann@31179: haftmann@51143: instantiation natural :: random haftmann@51143: begin haftmann@51143: haftmann@51143: definition random_natural :: "natural \ Random.seed haftmann@51143: \ (natural \ (unit \ Code_Evaluation.term)) \ Random.seed" haftmann@51143: where haftmann@51143: "random_natural i = Random.range (i + 1) \\ (\n. Pair (n, \_. Code_Evaluation.term_of n))" haftmann@51143: haftmann@51143: instance .. haftmann@51143: haftmann@51143: end haftmann@51143: haftmann@51143: instantiation integer :: random haftmann@51143: begin haftmann@51143: haftmann@51143: definition random_integer :: "natural \ Random.seed haftmann@51143: \ (integer \ (unit \ Code_Evaluation.term)) \ Random.seed" haftmann@51143: where haftmann@51143: "random_integer i = Random.range (2 * i + 1) \\ (\k. Pair ( haftmann@51143: let j = (if k \ i then integer_of_natural (k - i) else - (integer_of_natural (i - k))) haftmann@51143: in (j, \_. Code_Evaluation.term_of j)))" haftmann@51143: haftmann@51143: instance .. haftmann@51143: haftmann@51143: end haftmann@51143: haftmann@31260: haftmann@31260: subsection {* Complex generators *} haftmann@31260: haftmann@31603: text {* Towards @{typ "'a \ 'b"} *} haftmann@31603: haftmann@31603: axiomatization random_fun_aux :: "typerep \ typerep \ ('a \ 'a \ bool) \ ('a \ term) wenzelm@46975: \ (Random.seed \ ('b \ (unit \ term)) \ Random.seed) wenzelm@46975: \ (Random.seed \ Random.seed \ Random.seed) haftmann@31603: \ Random.seed \ (('a \ 'b) \ (unit \ term)) \ Random.seed" haftmann@31603: haftmann@31622: definition random_fun_lift :: "(Random.seed \ ('b \ (unit \ term)) \ Random.seed) wenzelm@46975: \ Random.seed \ (('a\term_of \ 'b\typerep) \ (unit \ term)) \ Random.seed" wenzelm@46975: where wenzelm@46975: "random_fun_lift f = wenzelm@46975: random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed" haftmann@31603: haftmann@38857: instantiation "fun" :: ("{equal, term_of}", random) random haftmann@31603: begin haftmann@31603: wenzelm@46975: definition haftmann@51143: random_fun :: "natural \ Random.seed \ (('a \ 'b) \ (unit \ term)) \ Random.seed" wenzelm@46975: where "random i = random_fun_lift (random i)" haftmann@31603: haftmann@31603: instance .. haftmann@31603: haftmann@31603: end haftmann@31603: haftmann@31603: text {* Towards type copies and datatypes *} haftmann@31603: wenzelm@46975: definition collapse :: "('a \ ('a \ 'b \ 'a) \ 'a) \ 'a \ 'b \ 'a" wenzelm@46975: where "collapse f = (f \\ id)" haftmann@31223: haftmann@51143: definition beyond :: "natural \ natural \ natural" wenzelm@46975: where "beyond k l = (if l > k then l else 0)" haftmann@31260: wenzelm@46975: lemma beyond_zero: "beyond k 0 = 0" haftmann@31267: by (simp add: beyond_def) haftmann@31267: bulwahn@46311: wenzelm@46975: definition (in term_syntax) [code_unfold]: wenzelm@46975: "valterm_emptyset = Code_Evaluation.valtermify ({} :: ('a :: typerep) set)" wenzelm@46975: wenzelm@46975: definition (in term_syntax) [code_unfold]: wenzelm@46975: "valtermify_insert x s = Code_Evaluation.valtermify insert {\} (x :: ('a :: typerep * _)) {\} s" bulwahn@46311: bulwahn@46311: instantiation set :: (random) random bulwahn@46311: begin bulwahn@46311: bulwahn@46311: primrec random_aux_set bulwahn@46311: where bulwahn@46311: "random_aux_set 0 j = collapse (Random.select_weight [(1, Pair valterm_emptyset)])" wenzelm@46975: | "random_aux_set (Code_Numeral.Suc i) j = wenzelm@46975: collapse (Random.select_weight wenzelm@46975: [(1, Pair valterm_emptyset), wenzelm@46975: (Code_Numeral.Suc i, wenzelm@46975: random j \\ (%x. random_aux_set i j \\ (%s. Pair (valtermify_insert x s))))])" bulwahn@46311: bulwahn@46311: lemma [code]: wenzelm@46975: "random_aux_set i j = wenzelm@46975: collapse (Random.select_weight [(1, Pair valterm_emptyset), wenzelm@46975: (i, random j \\ (%x. random_aux_set (i - 1) j \\ (%s. Pair (valtermify_insert x s))))])" haftmann@51143: proof (induct i rule: natural.induct) bulwahn@46311: case zero haftmann@50046: show ?case by (subst select_weight_drop_zero [symmetric]) haftmann@51143: (simp add: random_aux_set.simps [simplified] less_natural_def) bulwahn@46311: next huffman@46547: case (Suc i) haftmann@51143: show ?case by (simp only: random_aux_set.simps(2) [of "i"] Suc_natural_minus_one) bulwahn@46311: qed bulwahn@46311: wenzelm@46975: definition "random_set i = random_aux_set i i" bulwahn@46311: bulwahn@46311: instance .. bulwahn@46311: bulwahn@46311: end bulwahn@46311: haftmann@31483: lemma random_aux_rec: haftmann@51143: fixes random_aux :: "natural \ 'a" haftmann@31483: assumes "random_aux 0 = rhs 0" huffman@46547: and "\k. random_aux (Code_Numeral.Suc k) = rhs (Code_Numeral.Suc k)" haftmann@31483: shows "random_aux k = rhs k" haftmann@51143: using assms by (rule natural.induct) haftmann@31483: bulwahn@45718: subsection {* Deriving random generators for datatypes *} bulwahn@45718: wenzelm@48891: ML_file "Tools/Quickcheck/quickcheck_common.ML" wenzelm@48891: ML_file "Tools/Quickcheck/random_generators.ML" bulwahn@41923: setup Random_Generators.setup haftmann@34968: haftmann@34968: haftmann@34968: subsection {* Code setup *} blanchet@33561: haftmann@52435: code_printing haftmann@52435: constant random_fun_aux \ (Quickcheck) "Random'_Generators.random'_fun" haftmann@34968: -- {* With enough criminal energy this can be abused to derive @{prop False}; haftmann@34968: for this reason we use a distinguished target @{text Quickcheck} haftmann@34968: not spoiling the regular trusted code generation *} haftmann@34968: bulwahn@41935: code_reserved Quickcheck Random_Generators haftmann@34968: haftmann@37751: no_notation fcomp (infixl "\>" 60) haftmann@37751: no_notation scomp (infixl "\\" 60) haftmann@51126: bulwahn@45801: hide_const (open) catch_match random collapse beyond random_fun_aux random_fun_lift bulwahn@45801: haftmann@51126: hide_fact (open) collapse_def beyond_def random_fun_lift_def haftmann@31267: haftmann@31179: end haftmann@49972: