src/HOL/Nitpick.thy
changeset 57992 2371bff894f9
parent 57242 25aff3b8d550
child 58152 6fe60a9a5bad
equal deleted inserted replaced
57991:f50b3726266f 57992:2371bff894f9
    12 keywords
    12 keywords
    13   "nitpick" :: diag and
    13   "nitpick" :: diag and
    14   "nitpick_params" :: thy_decl
    14   "nitpick_params" :: thy_decl
    15 begin
    15 begin
    16 
    16 
       
    17 datatype ('a, 'b) fun_box = FunBox "'a \<Rightarrow> 'b"
       
    18 datatype ('a, 'b) pair_box = PairBox 'a 'b
       
    19 datatype 'a word = Word "'a set"
       
    20 
    17 typedecl bisim_iterator
    21 typedecl bisim_iterator
    18 
       
    19 axiomatization unknown :: 'a
       
    20            and is_unknown :: "'a \<Rightarrow> bool"
       
    21            and bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
       
    22            and bisim_iterator_max :: bisim_iterator
       
    23            and Quot :: "'a \<Rightarrow> 'b"
       
    24            and safe_The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
       
    25 
       
    26 datatype ('a, 'b) fun_box = FunBox "('a \<Rightarrow> 'b)"
       
    27 datatype ('a, 'b) pair_box = PairBox 'a 'b
       
    28 
       
    29 typedecl unsigned_bit
    22 typedecl unsigned_bit
    30 typedecl signed_bit
    23 typedecl signed_bit
    31 
    24 
    32 datatype 'a word = Word "('a set)"
    25 consts
       
    26   unknown :: 'a
       
    27   is_unknown :: "'a \<Rightarrow> bool"
       
    28   bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
       
    29   bisim_iterator_max :: bisim_iterator
       
    30   Quot :: "'a \<Rightarrow> 'b"
       
    31   safe_The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
    33 
    32 
    34 text {*
    33 text {*
    35 Alternative definitions.
    34 Alternative definitions.
    36 *}
    35 *}
    37 
    36 
    38 lemma Ex1_unfold [nitpick_unfold]:
    37 lemma Ex1_unfold[nitpick_unfold]: "Ex1 P \<equiv> \<exists>x. {x. P x} = {x}"
    39 "Ex1 P \<equiv> \<exists>x. {x. P x} = {x}"
    38   apply (rule eq_reflection)
    40 apply (rule eq_reflection)
    39   apply (simp add: Ex1_def set_eq_iff)
    41 apply (simp add: Ex1_def set_eq_iff)
    40   apply (rule iffI)
    42 apply (rule iffI)
    41    apply (erule exE)
    43  apply (erule exE)
    42    apply (erule conjE)
    44  apply (erule conjE)
    43    apply (rule_tac x = x in exI)
    45  apply (rule_tac x = x in exI)
    44    apply (rule allI)
    46  apply (rule allI)
    45    apply (rename_tac y)
    47  apply (rename_tac y)
    46    apply (erule_tac x = y in allE)
    48  apply (erule_tac x = y in allE)
    47   by auto
    49 by auto
    48 
    50 
    49 lemma rtrancl_unfold[nitpick_unfold]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
    51 lemma rtrancl_unfold [nitpick_unfold]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
       
    52   by (simp only: rtrancl_trancl_reflcl)
    50   by (simp only: rtrancl_trancl_reflcl)
    53 
    51 
    54 lemma rtranclp_unfold [nitpick_unfold]:
    52 lemma rtranclp_unfold[nitpick_unfold]: "rtranclp r a b \<equiv> (a = b \<or> tranclp r a b)"
    55 "rtranclp r a b \<equiv> (a = b \<or> tranclp r a b)"
    53   by (rule eq_reflection) (auto dest: rtranclpD)
    56 by (rule eq_reflection) (auto dest: rtranclpD)
    54 
    57 
    55 lemma tranclp_unfold[nitpick_unfold]:
    58 lemma tranclp_unfold [nitpick_unfold]:
    56   "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
    59 "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
    57   by (simp add: trancl_def)
    60 by (simp add: trancl_def)
       
    61 
    58 
    62 lemma [nitpick_simp]:
    59 lemma [nitpick_simp]:
    63 "of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))"
    60   "of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))"
    64 by (cases n) auto
    61   by (cases n) auto
    65 
    62 
    66 definition prod :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
    63 definition prod :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
    67 "prod A B = {(a, b). a \<in> A \<and> b \<in> B}"
    64   "prod A B = {(a, b). a \<in> A \<and> b \<in> B}"
    68 
    65 
    69 definition refl' :: "('a \<times> 'a) set \<Rightarrow> bool" where
    66 definition refl' :: "('a \<times> 'a) set \<Rightarrow> bool" where
    70 "refl' r \<equiv> \<forall>x. (x, x) \<in> r"
    67   "refl' r \<equiv> \<forall>x. (x, x) \<in> r"
    71 
    68 
    72 definition wf' :: "('a \<times> 'a) set \<Rightarrow> bool" where
    69 definition wf' :: "('a \<times> 'a) set \<Rightarrow> bool" where
    73 "wf' r \<equiv> acyclic r \<and> (finite r \<or> unknown)"
    70   "wf' r \<equiv> acyclic r \<and> (finite r \<or> unknown)"
    74 
    71 
    75 definition card' :: "'a set \<Rightarrow> nat" where
    72 definition card' :: "'a set \<Rightarrow> nat" where
    76 "card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
    73   "card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
    77 
    74 
    78 definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
    75 definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
    79 "setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
    76   "setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
    80 
    77 
    81 inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool" where
    78 inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool" where
    82 "fold_graph' f z {} z" |
    79   "fold_graph' f z {} z" |
    83 "\<lbrakk>x \<in> A; fold_graph' f z (A - {x}) y\<rbrakk> \<Longrightarrow> fold_graph' f z A (f x y)"
    80   "\<lbrakk>x \<in> A; fold_graph' f z (A - {x}) y\<rbrakk> \<Longrightarrow> fold_graph' f z A (f x y)"
    84 
    81 
    85 text {*
    82 text {*
    86 The following lemmas are not strictly necessary but they help the
    83 The following lemmas are not strictly necessary but they help the
    87 \textit{specialize} optimization.
    84 \textit{specialize} optimization.
    88 *}
    85 *}
    89 
    86 
    90 lemma The_psimp [nitpick_psimp]:
    87 lemma The_psimp[nitpick_psimp]: "P = (op =) x \<Longrightarrow> The P = x"
    91   "P = (op =) x \<Longrightarrow> The P = x"
       
    92   by auto
    88   by auto
    93 
    89 
    94 lemma Eps_psimp [nitpick_psimp]:
    90 lemma Eps_psimp[nitpick_psimp]:
    95 "\<lbrakk>P x; \<not> P y; Eps P = y\<rbrakk> \<Longrightarrow> Eps P = x"
    91   "\<lbrakk>P x; \<not> P y; Eps P = y\<rbrakk> \<Longrightarrow> Eps P = x"
    96 apply (cases "P (Eps P)")
    92   apply (cases "P (Eps P)")
    97  apply auto
    93    apply auto
    98 apply (erule contrapos_np)
    94   apply (erule contrapos_np)
    99 by (rule someI)
    95   by (rule someI)
   100 
    96 
   101 lemma case_unit_unfold [nitpick_unfold]:
    97 lemma case_unit_unfold[nitpick_unfold]:
   102 "case_unit x u \<equiv> x"
    98   "case_unit x u \<equiv> x"
   103 apply (subgoal_tac "u = ()")
    99   apply (subgoal_tac "u = ()")
   104  apply (simp only: unit.case)
   100    apply (simp only: unit.case)
   105 by simp
   101   by simp
   106 
   102 
   107 declare unit.case [nitpick_simp del]
   103 declare unit.case[nitpick_simp del]
   108 
   104 
   109 lemma case_nat_unfold [nitpick_unfold]:
   105 lemma case_nat_unfold[nitpick_unfold]:
   110 "case_nat x f n \<equiv> if n = 0 then x else f (n - 1)"
   106   "case_nat x f n \<equiv> if n = 0 then x else f (n - 1)"
   111 apply (rule eq_reflection)
   107   apply (rule eq_reflection)
   112 by (cases n) auto
   108   by (cases n) auto
   113 
   109 
   114 declare nat.case [nitpick_simp del]
   110 declare nat.case[nitpick_simp del]
   115 
   111 
   116 lemma size_list_simp [nitpick_simp]:
   112 lemma size_list_simp[nitpick_simp]:
   117 "size_list f xs = (if xs = [] then 0 else Suc (f (hd xs) + size_list f (tl xs)))"
   113   "size_list f xs = (if xs = [] then 0 else Suc (f (hd xs) + size_list f (tl xs)))"
   118 "size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
   114   "size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
   119 by (cases xs) auto
   115   by (cases xs) auto
   120 
   116 
   121 text {*
   117 text {*
   122 Auxiliary definitions used to provide an alternative representation for
   118 Auxiliary definitions used to provide an alternative representation for
   123 @{text rat} and @{text real}.
   119 @{text rat} and @{text real}.
   124 *}
   120 *}
   125 
   121 
   126 function nat_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
   122 function nat_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
   127 [simp del]: "nat_gcd x y = (if y = 0 then x else nat_gcd y (x mod y))"
   123   "nat_gcd x y = (if y = 0 then x else nat_gcd y (x mod y))"
   128 by auto
   124   by auto
   129 termination
   125   termination
   130 apply (relation "measure (\<lambda>(x, y). x + y + (if y > x then 1 else 0))")
   126   apply (relation "measure (\<lambda>(x, y). x + y + (if y > x then 1 else 0))")
   131  apply auto
   127    apply auto
   132  apply (metis mod_less_divisor xt1(9))
   128    apply (metis mod_less_divisor xt1(9))
   133 by (metis mod_mod_trivial mod_self nat_neq_iff xt1(10))
   129   by (metis mod_mod_trivial mod_self nat_neq_iff xt1(10))
       
   130 
       
   131 declare nat_gcd.simps[simp del]
   134 
   132 
   135 definition nat_lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
   133 definition nat_lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
   136 "nat_lcm x y = x * y div (nat_gcd x y)"
   134   "nat_lcm x y = x * y div (nat_gcd x y)"
   137 
   135 
   138 definition int_gcd :: "int \<Rightarrow> int \<Rightarrow> int" where
   136 definition int_gcd :: "int \<Rightarrow> int \<Rightarrow> int" where
   139 "int_gcd x y = int (nat_gcd (nat (abs x)) (nat (abs y)))"
   137   "int_gcd x y = int (nat_gcd (nat (abs x)) (nat (abs y)))"
   140 
   138 
   141 definition int_lcm :: "int \<Rightarrow> int \<Rightarrow> int" where
   139 definition int_lcm :: "int \<Rightarrow> int \<Rightarrow> int" where
   142 "int_lcm x y = int (nat_lcm (nat (abs x)) (nat (abs y)))"
   140   "int_lcm x y = int (nat_lcm (nat (abs x)) (nat (abs y)))"
   143 
   141 
   144 definition Frac :: "int \<times> int \<Rightarrow> bool" where
   142 definition Frac :: "int \<times> int \<Rightarrow> bool" where
   145 "Frac \<equiv> \<lambda>(a, b). b > 0 \<and> int_gcd a b = 1"
   143   "Frac \<equiv> \<lambda>(a, b). b > 0 \<and> int_gcd a b = 1"
   146 
   144 
   147 axiomatization
   145 consts
   148   Abs_Frac :: "int \<times> int \<Rightarrow> 'a" and
   146   Abs_Frac :: "int \<times> int \<Rightarrow> 'a"
   149   Rep_Frac :: "'a \<Rightarrow> int \<times> int"
   147   Rep_Frac :: "'a \<Rightarrow> int \<times> int"
   150 
   148 
   151 definition zero_frac :: 'a where
   149 definition zero_frac :: 'a where
   152 "zero_frac \<equiv> Abs_Frac (0, 1)"
   150   "zero_frac \<equiv> Abs_Frac (0, 1)"
   153 
   151 
   154 definition one_frac :: 'a where
   152 definition one_frac :: 'a where
   155 "one_frac \<equiv> Abs_Frac (1, 1)"
   153   "one_frac \<equiv> Abs_Frac (1, 1)"
   156 
   154 
   157 definition num :: "'a \<Rightarrow> int" where
   155 definition num :: "'a \<Rightarrow> int" where
   158 "num \<equiv> fst o Rep_Frac"
   156   "num \<equiv> fst o Rep_Frac"
   159 
   157 
   160 definition denom :: "'a \<Rightarrow> int" where
   158 definition denom :: "'a \<Rightarrow> int" where
   161 "denom \<equiv> snd o Rep_Frac"
   159   "denom \<equiv> snd o Rep_Frac"
   162 
   160 
   163 function norm_frac :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
   161 function norm_frac :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
   164 [simp del]: "norm_frac a b = (if b < 0 then norm_frac (- a) (- b)
   162   "norm_frac a b =
   165                               else if a = 0 \<or> b = 0 then (0, 1)
   163     (if b < 0 then norm_frac (- a) (- b)
   166                               else let c = int_gcd a b in (a div c, b div c))"
   164      else if a = 0 \<or> b = 0 then (0, 1)
   167 by pat_completeness auto
   165      else let c = int_gcd a b in (a div c, b div c))"
   168 termination by (relation "measure (\<lambda>(_, b). if b < 0 then 1 else 0)") auto
   166   by pat_completeness auto
       
   167   termination by (relation "measure (\<lambda>(_, b). if b < 0 then 1 else 0)") auto
       
   168 
       
   169 declare norm_frac.simps[simp del]
   169 
   170 
   170 definition frac :: "int \<Rightarrow> int \<Rightarrow> 'a" where
   171 definition frac :: "int \<Rightarrow> int \<Rightarrow> 'a" where
   171 "frac a b \<equiv> Abs_Frac (norm_frac a b)"
   172   "frac a b \<equiv> Abs_Frac (norm_frac a b)"
   172 
   173 
   173 definition plus_frac :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   174 definition plus_frac :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   174 [nitpick_simp]:
   175   [nitpick_simp]: "plus_frac q r = (let d = int_lcm (denom q) (denom r) in
   175 "plus_frac q r = (let d = int_lcm (denom q) (denom r) in
   176     frac (num q * (d div denom q) + num r * (d div denom r)) d)"
   176                     frac (num q * (d div denom q) + num r * (d div denom r)) d)"
       
   177 
   177 
   178 definition times_frac :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   178 definition times_frac :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   179 [nitpick_simp]:
   179   [nitpick_simp]: "times_frac q r = frac (num q * num r) (denom q * denom r)"
   180 "times_frac q r = frac (num q * num r) (denom q * denom r)"
       
   181 
   180 
   182 definition uminus_frac :: "'a \<Rightarrow> 'a" where
   181 definition uminus_frac :: "'a \<Rightarrow> 'a" where
   183 "uminus_frac q \<equiv> Abs_Frac (- num q, denom q)"
   182   "uminus_frac q \<equiv> Abs_Frac (- num q, denom q)"
   184 
   183 
   185 definition number_of_frac :: "int \<Rightarrow> 'a" where
   184 definition number_of_frac :: "int \<Rightarrow> 'a" where
   186 "number_of_frac n \<equiv> Abs_Frac (n, 1)"
   185   "number_of_frac n \<equiv> Abs_Frac (n, 1)"
   187 
   186 
   188 definition inverse_frac :: "'a \<Rightarrow> 'a" where
   187 definition inverse_frac :: "'a \<Rightarrow> 'a" where
   189 "inverse_frac q \<equiv> frac (denom q) (num q)"
   188   "inverse_frac q \<equiv> frac (denom q) (num q)"
   190 
   189 
   191 definition less_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
   190 definition less_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
   192 [nitpick_simp]:
   191   [nitpick_simp]: "less_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) < 0"
   193 "less_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) < 0"
       
   194 
   192 
   195 definition less_eq_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
   193 definition less_eq_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
   196 [nitpick_simp]:
   194   [nitpick_simp]: "less_eq_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) \<le> 0"
   197 "less_eq_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) \<le> 0"
       
   198 
   195 
   199 definition of_frac :: "'a \<Rightarrow> 'b\<Colon>{inverse,ring_1}" where
   196 definition of_frac :: "'a \<Rightarrow> 'b\<Colon>{inverse,ring_1}" where
   200 "of_frac q \<equiv> of_int (num q) / of_int (denom q)"
   197   "of_frac q \<equiv> of_int (num q) / of_int (denom q)"
   201 
   198 
   202 axiomatization wf_wfrec :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
   199 axiomatization wf_wfrec :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
   203 
   200 
   204 definition wf_wfrec' :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
   201 definition wf_wfrec' :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
   205 [nitpick_simp]: "wf_wfrec' R F x = F (cut (wf_wfrec R F) R x) x"
   202   [nitpick_simp]: "wf_wfrec' R F x = F (cut (wf_wfrec R F) R x) x"
   206 
   203 
   207 definition wfrec' ::  "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
   204 definition wfrec' ::  "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
   208 "wfrec' R F x \<equiv> if wf R then wf_wfrec' R F x
   205   "wfrec' R F x \<equiv> if wf R then wf_wfrec' R F x else THE y. wfrec_rel R (\<lambda>f x. F (cut f R x) x) x y"
   209                 else THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
       
   210 
   206 
   211 ML_file "Tools/Nitpick/kodkod.ML"
   207 ML_file "Tools/Nitpick/kodkod.ML"
   212 ML_file "Tools/Nitpick/kodkod_sat.ML"
   208 ML_file "Tools/Nitpick/kodkod_sat.ML"
   213 ML_file "Tools/Nitpick/nitpick_util.ML"
   209 ML_file "Tools/Nitpick/nitpick_util.ML"
   214 ML_file "Tools/Nitpick/nitpick_hol.ML"
   210 ML_file "Tools/Nitpick/nitpick_hol.ML"
   232      (@{const_name wf}, @{const_name wf'}),
   228      (@{const_name wf}, @{const_name wf'}),
   233      (@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
   229      (@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
   234      (@{const_name wfrec}, @{const_name wfrec'})]
   230      (@{const_name wfrec}, @{const_name wfrec'})]
   235 *}
   231 *}
   236 
   232 
   237 hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The
   233 hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The FunBox PairBox Word prod
   238     FunBox PairBox Word prod refl' wf' card' setsum'
   234   refl' wf' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac
   239     fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac zero_frac
   235   zero_frac one_frac num denom norm_frac frac plus_frac times_frac uminus_frac number_of_frac
   240     one_frac num denom norm_frac frac plus_frac times_frac uminus_frac
   236   inverse_frac less_frac less_eq_frac of_frac wf_wfrec wf_wfrec wfrec'
   241     number_of_frac inverse_frac less_frac less_eq_frac of_frac wf_wfrec wf_wfrec
   237 
   242     wfrec'
       
   243 hide_type (open) bisim_iterator fun_box pair_box unsigned_bit signed_bit word
   238 hide_type (open) bisim_iterator fun_box pair_box unsigned_bit signed_bit word
   244 hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold
   239 
   245     prod_def refl'_def wf'_def card'_def setsum'_def
   240 hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold prod_def refl'_def wf'_def
   246     fold_graph'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold
   241   card'_def setsum'_def fold_graph'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold
   247     size_list_simp nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def
   242   size_list_simp nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def zero_frac_def one_frac_def
   248     zero_frac_def one_frac_def num_def denom_def norm_frac_def frac_def
   243   num_def denom_def norm_frac_def frac_def plus_frac_def times_frac_def uminus_frac_def
   249     plus_frac_def times_frac_def uminus_frac_def number_of_frac_def
   244   number_of_frac_def inverse_frac_def less_frac_def less_eq_frac_def of_frac_def wf_wfrec'_def
   250     inverse_frac_def less_frac_def less_eq_frac_def of_frac_def wf_wfrec'_def
   245   wfrec'_def
   251     wfrec'_def
       
   252 
   246 
   253 end
   247 end