(* Title: HOL/Nitpick.thy Author: Jasmin Blanchette, TU Muenchen Copyright 2008, 2009, 2010 Nitpick: Yet another counterexample generator for Isabelle/HOL. *) header {* Nitpick: Yet Another Counterexample Generator for Isabelle/HOL *} theory Nitpick imports Map Quotient SAT Record uses ("Tools/Nitpick/kodkod.ML") ("Tools/Nitpick/kodkod_sat.ML") ("Tools/Nitpick/nitpick_util.ML") ("Tools/Nitpick/nitpick_hol.ML") ("Tools/Nitpick/nitpick_preproc.ML") ("Tools/Nitpick/nitpick_mono.ML") ("Tools/Nitpick/nitpick_scope.ML") ("Tools/Nitpick/nitpick_peephole.ML") ("Tools/Nitpick/nitpick_rep.ML") ("Tools/Nitpick/nitpick_nut.ML") ("Tools/Nitpick/nitpick_kodkod.ML") ("Tools/Nitpick/nitpick_model.ML") ("Tools/Nitpick/nitpick.ML") ("Tools/Nitpick/nitpick_isar.ML") ("Tools/Nitpick/nitpick_tests.ML") ("Tools/Nitpick/minipick.ML") begin typedecl bisim_iterator axiomatization unknown :: 'a and is_unknown :: "'a \ bool" and undefined_fast_The :: 'a and undefined_fast_Eps :: 'a and bisim :: "bisim_iterator \ 'a \ 'a \ bool" and bisim_iterator_max :: bisim_iterator and Quot :: "'a \ 'b" and safe_The :: "('a \ bool) \ 'a" and safe_Eps :: "('a \ bool) \ 'a" datatype ('a, 'b) fin_fun = FinFun "('a \ 'b)" datatype ('a, 'b) fun_box = FunBox "('a \ 'b)" datatype ('a, 'b) pair_box = PairBox 'a 'b typedecl unsigned_bit typedecl signed_bit datatype 'a word = Word "('a set)" text {* Alternative definitions. *} lemma If_def [nitpick_def, no_atp]: "(if P then Q else R) \ (P \ Q) \ (\ P \ R)" by (rule eq_reflection) (rule if_bool_eq_conj) lemma Ex1_def [nitpick_def, no_atp]: "Ex1 P \ \x. P = {x}" apply (rule eq_reflection) apply (simp add: Ex1_def expand_set_eq) apply (rule iffI) apply (erule exE) apply (erule conjE) apply (rule_tac x = x in exI) apply (rule allI) apply (rename_tac y) apply (erule_tac x = y in allE) by (auto simp: mem_def) lemma rtrancl_def [nitpick_def, no_atp]: "r\<^sup>* \ (r\<^sup>+)\<^sup>=" by simp lemma rtranclp_def [nitpick_def, no_atp]: "rtranclp r a b \ (a = b \ tranclp r a b)" by (rule eq_reflection) (auto dest: rtranclpD) lemma tranclp_def [nitpick_def, no_atp]: "tranclp r a b \ trancl (split r) (a, b)" by (simp add: trancl_def Collect_def mem_def) definition refl' :: "('a \ 'a \ bool) \ bool" where "refl' r \ \x. (x, x) \ r" definition wf' :: "('a \ 'a \ bool) \ bool" where "wf' r \ acyclic r \ (finite r \ unknown)" axiomatization wf_wfrec :: "('a \ 'a \ bool) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" definition wf_wfrec' :: "('a \ 'a \ bool) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" where [nitpick_simp]: "wf_wfrec' R F x = F (Recdef.cut (wf_wfrec R F) R x) x" definition wfrec' :: "('a \ 'a \ bool) \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" where "wfrec' R F x \ if wf R then wf_wfrec' R F x else THE y. wfrec_rel R (%f x. F (Recdef.cut f R x) x) x y" definition card' :: "('a \ bool) \ nat" where "card' A \ if finite A then length (safe_Eps (\xs. set xs = A \ distinct xs)) else 0" definition setsum' :: "('a \ 'b\comm_monoid_add) \ ('a \ bool) \ 'b" where "setsum' f A \ if finite A then listsum (map f (safe_Eps (\xs. set xs = A \ distinct xs))) else 0" inductive fold_graph' :: "('a \ 'b \ 'b) \ 'b \ ('a \ bool) \ 'b \ bool" where "fold_graph' f z {} z" | "\x \ A; fold_graph' f z (A - {x}) y\ \ fold_graph' f z A (f x y)" text {* The following lemmas are not strictly necessary but they help the \textit{special\_level} optimization. *} lemma The_psimp [nitpick_psimp, no_atp]: "P = {x} \ The P = x" by (subgoal_tac "{x} = (\y. y = x)") (auto simp: mem_def) lemma Eps_psimp [nitpick_psimp, no_atp]: "\P x; \ P y; Eps P = y\ \ Eps P = x" apply (case_tac "P (Eps P)") apply auto apply (erule contrapos_np) by (rule someI) lemma unit_case_def [nitpick_def, no_atp]: "unit_case x u \ x" apply (subgoal_tac "u = ()") apply (simp only: unit.cases) by simp declare unit.cases [nitpick_simp del] lemma nat_case_def [nitpick_def, no_atp]: "nat_case x f n \ if n = 0 then x else f (n - 1)" apply (rule eq_reflection) by (case_tac n) auto declare nat.cases [nitpick_simp del] lemma list_size_simp [nitpick_simp, no_atp]: "list_size f xs = (if xs = [] then 0 else Suc (f (hd xs) + list_size f (tl xs)))" "size xs = (if xs = [] then 0 else Suc (size (tl xs)))" by (case_tac xs) auto text {* Auxiliary definitions used to provide an alternative representation for @{text rat} and @{text real}. *} function nat_gcd :: "nat \ nat \ nat" where [simp del]: "nat_gcd x y = (if y = 0 then x else nat_gcd y (x mod y))" by auto termination apply (relation "measure (\(x, y). x + y + (if y > x then 1 else 0))") apply auto apply (metis mod_less_divisor xt1(9)) by (metis mod_mod_trivial mod_self nat_neq_iff xt1(10)) definition nat_lcm :: "nat \ nat \ nat" where "nat_lcm x y = x * y div (nat_gcd x y)" definition int_gcd :: "int \ int \ int" where "int_gcd x y = int (nat_gcd (nat (abs x)) (nat (abs y)))" definition int_lcm :: "int \ int \ int" where "int_lcm x y = int (nat_lcm (nat (abs x)) (nat (abs y)))" definition Frac :: "int \ int \ bool" where "Frac \ \(a, b). b > 0 \ int_gcd a b = 1" axiomatization Abs_Frac :: "int \ int \ 'a" and Rep_Frac :: "'a \ int \ int" definition zero_frac :: 'a where "zero_frac \ Abs_Frac (0, 1)" definition one_frac :: 'a where "one_frac \ Abs_Frac (1, 1)" definition num :: "'a \ int" where "num \ fst o Rep_Frac" definition denom :: "'a \ int" where "denom \ snd o Rep_Frac" function norm_frac :: "int \ int \ int \ int" where [simp del]: "norm_frac a b = (if b < 0 then norm_frac (- a) (- b) else if a = 0 \ b = 0 then (0, 1) else let c = int_gcd a b in (a div c, b div c))" by pat_completeness auto termination by (relation "measure (\(_, b). if b < 0 then 1 else 0)") auto definition frac :: "int \ int \ 'a" where "frac a b \ Abs_Frac (norm_frac a b)" definition plus_frac :: "'a \ 'a \ 'a" where [nitpick_simp]: "plus_frac q r = (let d = int_lcm (denom q) (denom r) in frac (num q * (d div denom q) + num r * (d div denom r)) d)" definition times_frac :: "'a \ 'a \ 'a" where [nitpick_simp]: "times_frac q r = frac (num q * num r) (denom q * denom r)" definition uminus_frac :: "'a \ 'a" where "uminus_frac q \ Abs_Frac (- num q, denom q)" definition number_of_frac :: "int \ 'a" where "number_of_frac n \ Abs_Frac (n, 1)" definition inverse_frac :: "'a \ 'a" where "inverse_frac q \ frac (denom q) (num q)" definition less_frac :: "'a \ 'a \ bool" where [nitpick_simp]: "less_frac q r \ num (plus_frac q (uminus_frac r)) < 0" definition less_eq_frac :: "'a \ 'a \ bool" where [nitpick_simp]: "less_eq_frac q r \ num (plus_frac q (uminus_frac r)) \ 0" definition of_frac :: "'a \ 'b\{inverse,ring_1}" where "of_frac q \ of_int (num q) / of_int (denom q)" use "Tools/Nitpick/kodkod.ML" use "Tools/Nitpick/kodkod_sat.ML" use "Tools/Nitpick/nitpick_util.ML" use "Tools/Nitpick/nitpick_hol.ML" use "Tools/Nitpick/nitpick_mono.ML" use "Tools/Nitpick/nitpick_preproc.ML" use "Tools/Nitpick/nitpick_scope.ML" use "Tools/Nitpick/nitpick_peephole.ML" use "Tools/Nitpick/nitpick_rep.ML" use "Tools/Nitpick/nitpick_nut.ML" use "Tools/Nitpick/nitpick_kodkod.ML" use "Tools/Nitpick/nitpick_model.ML" use "Tools/Nitpick/nitpick.ML" use "Tools/Nitpick/nitpick_isar.ML" use "Tools/Nitpick/nitpick_tests.ML" use "Tools/Nitpick/minipick.ML" setup {* Nitpick_Isar.setup *} hide_const (open) unknown is_unknown undefined_fast_The undefined_fast_Eps bisim bisim_iterator_max Quot safe_The safe_Eps FinFun FunBox PairBox Word refl' wf' wf_wfrec wf_wfrec' wfrec' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac zero_frac one_frac num denom norm_frac frac plus_frac times_frac uminus_frac number_of_frac inverse_frac less_frac less_eq_frac of_frac hide_type (open) bisim_iterator fin_fun fun_box pair_box unsigned_bit signed_bit word hide_fact (open) If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def wf'_def wf_wfrec'_def wfrec'_def card'_def setsum'_def fold_graph'_def The_psimp Eps_psimp unit_case_def nat_case_def list_size_simp nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def zero_frac_def one_frac_def num_def denom_def norm_frac_def frac_def plus_frac_def times_frac_def uminus_frac_def number_of_frac_def inverse_frac_def less_frac_def less_eq_frac_def of_frac_def end