# HG changeset patch # User blanchet # Date 1408433681 -7200 # Node ID 2371bff894f97dd075de003600347551b8f7e339 # Parent f50b3726266f3fe4eaa88e9bb144797b0f5e18b6 tuning diff -r f50b3726266f -r 2371bff894f9 src/HOL/Nitpick.thy --- a/src/HOL/Nitpick.thy Tue Aug 19 09:34:30 2014 +0200 +++ b/src/HOL/Nitpick.thy Tue Aug 19 09:34:41 2014 +0200 @@ -14,109 +14,105 @@ "nitpick_params" :: thy_decl begin -typedecl bisim_iterator +datatype ('a, 'b) fun_box = FunBox "'a \ 'b" +datatype ('a, 'b) pair_box = PairBox 'a 'b +datatype 'a word = Word "'a set" -axiomatization unknown :: 'a - and is_unknown :: "'a \ bool" - and bisim :: "bisim_iterator \ 'a \ 'a \ bool" - and bisim_iterator_max :: bisim_iterator - and Quot :: "'a \ 'b" - and safe_The :: "('a \ bool) \ 'a" - -datatype ('a, 'b) fun_box = FunBox "('a \ 'b)" -datatype ('a, 'b) pair_box = PairBox 'a 'b - +typedecl bisim_iterator typedecl unsigned_bit typedecl signed_bit -datatype 'a word = Word "('a set)" +consts + unknown :: 'a + is_unknown :: "'a \ bool" + bisim :: "bisim_iterator \ 'a \ 'a \ bool" + bisim_iterator_max :: bisim_iterator + Quot :: "'a \ 'b" + safe_The :: "('a \ bool) \ 'a" text {* Alternative definitions. *} -lemma Ex1_unfold [nitpick_unfold]: -"Ex1 P \ \x. {x. P x} = {x}" -apply (rule eq_reflection) -apply (simp add: Ex1_def set_eq_iff) -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 +lemma Ex1_unfold[nitpick_unfold]: "Ex1 P \ \x. {x. P x} = {x}" + apply (rule eq_reflection) + apply (simp add: Ex1_def set_eq_iff) + 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 -lemma rtrancl_unfold [nitpick_unfold]: "r\<^sup>* \ (r\<^sup>+)\<^sup>=" +lemma rtrancl_unfold[nitpick_unfold]: "r\<^sup>* \ (r\<^sup>+)\<^sup>=" by (simp only: rtrancl_trancl_reflcl) -lemma rtranclp_unfold [nitpick_unfold]: -"rtranclp r a b \ (a = b \ tranclp r a b)" -by (rule eq_reflection) (auto dest: rtranclpD) +lemma rtranclp_unfold[nitpick_unfold]: "rtranclp r a b \ (a = b \ tranclp r a b)" + by (rule eq_reflection) (auto dest: rtranclpD) -lemma tranclp_unfold [nitpick_unfold]: -"tranclp r a b \ (a, b) \ trancl {(x, y). r x y}" -by (simp add: trancl_def) +lemma tranclp_unfold[nitpick_unfold]: + "tranclp r a b \ (a, b) \ trancl {(x, y). r x y}" + by (simp add: trancl_def) lemma [nitpick_simp]: -"of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))" -by (cases n) auto + "of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))" + by (cases n) auto definition prod :: "'a set \ 'b set \ ('a \ 'b) set" where -"prod A B = {(a, b). a \ A \ b \ B}" + "prod A B = {(a, b). a \ A \ b \ B}" definition refl' :: "('a \ 'a) set \ bool" where -"refl' r \ \x. (x, x) \ r" + "refl' r \ \x. (x, x) \ r" definition wf' :: "('a \ 'a) set \ bool" where -"wf' r \ acyclic r \ (finite r \ unknown)" + "wf' r \ acyclic r \ (finite r \ unknown)" definition card' :: "'a set \ nat" where -"card' A \ if finite A then length (SOME xs. set xs = A \ distinct xs) else 0" + "card' A \ if finite A then length (SOME xs. set xs = A \ distinct xs) else 0" definition setsum' :: "('a \ 'b\comm_monoid_add) \ 'a set \ 'b" where -"setsum' f A \ if finite A then listsum (map f (SOME xs. set xs = A \ distinct xs)) else 0" + "setsum' f A \ if finite A then listsum (map f (SOME xs. set xs = A \ distinct xs)) else 0" inductive fold_graph' :: "('a \ 'b \ 'b) \ 'b \ 'a set \ '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)" + "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{specialize} optimization. *} -lemma The_psimp [nitpick_psimp]: - "P = (op =) x \ The P = x" +lemma The_psimp[nitpick_psimp]: "P = (op =) x \ The P = x" by auto -lemma Eps_psimp [nitpick_psimp]: -"\P x; \ P y; Eps P = y\ \ Eps P = x" -apply (cases "P (Eps P)") - apply auto -apply (erule contrapos_np) -by (rule someI) +lemma Eps_psimp[nitpick_psimp]: + "\P x; \ P y; Eps P = y\ \ Eps P = x" + apply (cases "P (Eps P)") + apply auto + apply (erule contrapos_np) + by (rule someI) -lemma case_unit_unfold [nitpick_unfold]: -"case_unit x u \ x" -apply (subgoal_tac "u = ()") - apply (simp only: unit.case) -by simp +lemma case_unit_unfold[nitpick_unfold]: + "case_unit x u \ x" + apply (subgoal_tac "u = ()") + apply (simp only: unit.case) + by simp -declare unit.case [nitpick_simp del] +declare unit.case[nitpick_simp del] -lemma case_nat_unfold [nitpick_unfold]: -"case_nat x f n \ if n = 0 then x else f (n - 1)" -apply (rule eq_reflection) -by (cases n) auto +lemma case_nat_unfold[nitpick_unfold]: + "case_nat x f n \ if n = 0 then x else f (n - 1)" + apply (rule eq_reflection) + by (cases n) auto -declare nat.case [nitpick_simp del] +declare nat.case[nitpick_simp del] -lemma size_list_simp [nitpick_simp]: -"size_list f xs = (if xs = [] then 0 else Suc (f (hd xs) + size_list f (tl xs)))" -"size xs = (if xs = [] then 0 else Suc (size (tl xs)))" -by (cases xs) auto +lemma size_list_simp[nitpick_simp]: + "size_list f xs = (if xs = [] then 0 else Suc (f (hd xs) + size_list f (tl xs)))" + "size xs = (if xs = [] then 0 else Suc (size (tl xs)))" + by (cases xs) auto text {* Auxiliary definitions used to provide an alternative representation for @@ -124,89 +120,89 @@ *} 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)) + "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)) + +declare nat_gcd.simps[simp del] definition nat_lcm :: "nat \ nat \ nat" where -"nat_lcm x y = x * y div (nat_gcd x y)" + "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)))" + "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)))" + "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" + "Frac \ \(a, b). b > 0 \ int_gcd a b = 1" -axiomatization - Abs_Frac :: "int \ int \ 'a" and +consts + Abs_Frac :: "int \ int \ 'a" Rep_Frac :: "'a \ int \ int" definition zero_frac :: 'a where -"zero_frac \ Abs_Frac (0, 1)" + "zero_frac \ Abs_Frac (0, 1)" definition one_frac :: 'a where -"one_frac \ Abs_Frac (1, 1)" + "one_frac \ Abs_Frac (1, 1)" definition num :: "'a \ int" where -"num \ fst o Rep_Frac" + "num \ fst o Rep_Frac" definition denom :: "'a \ int" where -"denom \ snd o Rep_Frac" + "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 + "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 + +declare norm_frac.simps[simp del] definition frac :: "int \ int \ 'a" where -"frac a b \ Abs_Frac (norm_frac a b)" + "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)" + [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)" + [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)" + "uminus_frac q \ Abs_Frac (- num q, denom q)" definition number_of_frac :: "int \ 'a" where -"number_of_frac n \ Abs_Frac (n, 1)" + "number_of_frac n \ Abs_Frac (n, 1)" definition inverse_frac :: "'a \ 'a" where -"inverse_frac q \ frac (denom q) (num q)" + "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" + [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" + [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)" + "of_frac q \ of_int (num q) / of_int (denom q)" axiomatization wf_wfrec :: "('a \ 'a) set \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" definition wf_wfrec' :: "('a \ 'a) set \ (('a \ 'b) \ 'a \ 'b) \ 'a \ 'b" where -[nitpick_simp]: "wf_wfrec' R F x = F (cut (wf_wfrec R F) R x) x" + [nitpick_simp]: "wf_wfrec' R F x = F (cut (wf_wfrec R F) R x) x" definition wfrec' :: "('a \ 'a) set \ (('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 (cut f R x) x) x y" + "wfrec' R F x \ if wf R then wf_wfrec' R F x else THE y. wfrec_rel R (\f x. F (cut f R x) x) x y" ML_file "Tools/Nitpick/kodkod.ML" ML_file "Tools/Nitpick/kodkod_sat.ML" @@ -234,20 +230,18 @@ (@{const_name wfrec}, @{const_name wfrec'})] *} -hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The - FunBox PairBox Word prod refl' wf' 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 wf_wfrec wf_wfrec - wfrec' +hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The FunBox PairBox Word prod + refl' wf' 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 wf_wfrec wf_wfrec wfrec' + hide_type (open) bisim_iterator fun_box pair_box unsigned_bit signed_bit word -hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold - prod_def refl'_def wf'_def card'_def setsum'_def - fold_graph'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold - size_list_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 wf_wfrec'_def - wfrec'_def + +hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold prod_def refl'_def wf'_def + card'_def setsum'_def fold_graph'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold + size_list_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 wf_wfrec'_def + wfrec'_def end diff -r f50b3726266f -r 2371bff894f9 src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Tue Aug 19 09:34:30 2014 +0200 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Tue Aug 19 09:34:41 2014 +0200 @@ -974,7 +974,7 @@ fun zero_const T = Const (@{const_name zero_class.zero}, T) fun suc_const T = Const (@{const_name Suc}, T --> T) -fun uncached_data_type_constrs ({thy, ctxt, ...} : hol_context) (T as Type (s, _)) = +fun uncached_data_type_constrs ({ctxt, ...} : hol_context) (T as Type (s, _)) = if is_interpreted_type s then [] else