| 64962 |      1 | (*  Title:      HOL/Decision_Procs/Conversions.thy
 | 
|  |      2 |     Author:     Stefan Berghofer
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | theory Conversions
 | 
|  |      6 | imports Main
 | 
|  |      7 | begin
 | 
|  |      8 | 
 | 
| 64998 |      9 | ML \<open>
 | 
| 64962 |     10 | fun tactic_of_conv cv i st =
 | 
|  |     11 |   if i > Thm.nprems_of st then Seq.empty
 | 
|  |     12 |   else Seq.single (Conv.gconv_rule cv i st);
 | 
|  |     13 | 
 | 
|  |     14 | fun binop_conv cv cv' = Conv.combination_conv (Conv.arg_conv cv) cv';
 | 
| 64998 |     15 | \<close>
 | 
| 64962 |     16 | 
 | 
| 64998 |     17 | ML \<open>
 | 
| 64962 |     18 | fun err s ct =
 | 
|  |     19 |    error (s ^ ": " ^ Syntax.string_of_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct));
 | 
| 64998 |     20 | \<close>
 | 
| 64962 |     21 | 
 | 
|  |     22 | attribute_setup meta =
 | 
| 64998 |     23 |   \<open>Scan.succeed (fn (ctxt, th) => (NONE, SOME (mk_meta_eq th)))\<close>
 | 
|  |     24 |   \<open>convert equality to meta equality\<close>
 | 
| 64962 |     25 | 
 | 
| 64998 |     26 | ML \<open>
 | 
| 64962 |     27 | fun strip_app ct = ct |> Drule.strip_comb |>> Thm.term_of |>> dest_Const |>> fst;
 | 
|  |     28 | 
 | 
|  |     29 | fun inst cTs cts th =
 | 
|  |     30 |   Thm.instantiate' (map SOME cTs) (map SOME cts) th;
 | 
|  |     31 | 
 | 
|  |     32 | fun transitive' eq eq' = Thm.transitive eq (eq' (Thm.rhs_of eq));
 | 
|  |     33 | 
 | 
|  |     34 | fun type_of_eqn eqn = Thm.ctyp_of_cterm (Thm.dest_arg1 (Thm.cprop_of eqn));
 | 
|  |     35 | 
 | 
|  |     36 | fun cong1 conv ct =
 | 
|  |     37 |   Thm.combination (Thm.reflexive (Thm.dest_fun ct)) (conv (Thm.dest_arg ct));
 | 
|  |     38 | 
 | 
|  |     39 | fun cong1' conv' conv ct =
 | 
|  |     40 |   let val eqn = conv (Thm.dest_arg ct)
 | 
|  |     41 |   in
 | 
|  |     42 |     Thm.transitive
 | 
|  |     43 |       (Thm.combination (Thm.reflexive (Thm.dest_fun ct)) eqn)
 | 
|  |     44 |       (conv' (Thm.rhs_of eqn))
 | 
|  |     45 |   end;
 | 
|  |     46 | 
 | 
|  |     47 | fun cong2 conv1 conv2 ct =
 | 
|  |     48 |   Thm.combination
 | 
|  |     49 |     (Thm.combination
 | 
|  |     50 |        (Thm.reflexive (Thm.dest_fun2 ct))
 | 
|  |     51 |        (conv1 (Thm.dest_arg1 ct)))
 | 
|  |     52 |     (conv2 (Thm.dest_arg ct));
 | 
|  |     53 | 
 | 
|  |     54 | fun cong2' conv conv1 conv2 ct =
 | 
|  |     55 |   let
 | 
|  |     56 |     val eqn1 = conv1 (Thm.dest_arg1 ct);
 | 
|  |     57 |     val eqn2 = conv2 (Thm.dest_arg ct)
 | 
|  |     58 |   in
 | 
|  |     59 |     Thm.transitive
 | 
|  |     60 |       (Thm.combination
 | 
|  |     61 |          (Thm.combination (Thm.reflexive (Thm.dest_fun2 ct)) eqn1)
 | 
|  |     62 |          eqn2)
 | 
|  |     63 |       (conv (Thm.rhs_of eqn1) (Thm.rhs_of eqn2))
 | 
|  |     64 |   end;
 | 
|  |     65 | 
 | 
|  |     66 | fun cong2'' conv eqn1 eqn2 =
 | 
|  |     67 |   let val eqn3 = conv (Thm.rhs_of eqn1) (Thm.rhs_of eqn2)
 | 
|  |     68 |   in
 | 
|  |     69 |     Thm.transitive
 | 
|  |     70 |       (Thm.combination
 | 
|  |     71 |          (Thm.combination (Thm.reflexive (Thm.dest_fun2 (Thm.lhs_of eqn3))) eqn1)
 | 
|  |     72 |          eqn2)
 | 
|  |     73 |       eqn3
 | 
|  |     74 |   end;
 | 
|  |     75 | 
 | 
|  |     76 | fun args1 conv ct = conv (Thm.dest_arg ct);
 | 
|  |     77 | fun args2 conv ct = conv (Thm.dest_arg1 ct) (Thm.dest_arg ct);
 | 
| 64998 |     78 | \<close>
 | 
| 64962 |     79 | 
 | 
| 64998 |     80 | ML \<open>
 | 
| 64962 |     81 | fun strip_numeral ct = (case strip_app ct of
 | 
| 69597 |     82 |     (\<^const_name>\<open>uminus\<close>, [n]) => (case strip_app n of
 | 
|  |     83 |       (\<^const_name>\<open>numeral\<close>, [b]) => (\<^const_name>\<open>uminus\<close>, [b])
 | 
| 64962 |     84 |     | _ => ("", []))
 | 
|  |     85 |   | x => x);
 | 
| 64998 |     86 | \<close>
 | 
| 64962 |     87 | 
 | 
|  |     88 | lemma nat_minus1_eq: "nat (- 1) = 0"
 | 
|  |     89 |   by simp
 | 
|  |     90 | 
 | 
| 64998 |     91 | ML \<open>
 | 
| 64962 |     92 | fun nat_conv i = (case strip_app i of
 | 
| 69597 |     93 |     (\<^const_name>\<open>zero_class.zero\<close>, []) => @{thm nat_0 [meta]}
 | 
|  |     94 |   | (\<^const_name>\<open>one_class.one\<close>, []) => @{thm nat_one_as_int [meta, symmetric]}
 | 
|  |     95 |   | (\<^const_name>\<open>numeral\<close>, [b]) => inst [] [b] @{thm nat_numeral [meta]}
 | 
|  |     96 |   | (\<^const_name>\<open>uminus\<close>, [b]) => (case strip_app b of
 | 
|  |     97 |       (\<^const_name>\<open>one_class.one\<close>, []) => @{thm nat_minus1_eq [meta]}
 | 
|  |     98 |     | (\<^const_name>\<open>numeral\<close>, [b']) => inst [] [b'] @{thm nat_neg_numeral [meta]}));
 | 
| 64998 |     99 | \<close>
 | 
| 64962 |    100 | 
 | 
| 64998 |    101 | ML \<open>
 | 
| 64962 |    102 | fun add_num_conv b b' = (case (strip_app b, strip_app b') of
 | 
| 69597 |    103 |     ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    104 |       @{thm add_num_simps(1) [meta]}
 | 
| 69597 |    105 |   | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    106 |       inst [] [n] @{thm add_num_simps(2) [meta]}
 | 
| 69597 |    107 |   | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    108 |       transitive'
 | 
|  |    109 |         (inst [] [n] @{thm add_num_simps(3) [meta]})
 | 
|  |    110 |         (cong1 (args2 add_num_conv))
 | 
| 69597 |    111 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    112 |       inst [] [m] @{thm add_num_simps(4) [meta]}
 | 
| 69597 |    113 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    114 |       transitive'
 | 
|  |    115 |         (inst [] [m, n] @{thm add_num_simps(5) [meta]})
 | 
|  |    116 |         (cong1 (args2 add_num_conv))
 | 
| 69597 |    117 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    118 |       transitive'
 | 
|  |    119 |         (inst [] [m, n] @{thm add_num_simps(6) [meta]})
 | 
|  |    120 |         (cong1 (args2 add_num_conv))
 | 
| 69597 |    121 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    122 |       transitive'
 | 
|  |    123 |         (inst [] [m] @{thm add_num_simps(7) [meta]})
 | 
|  |    124 |         (cong1 (args2 add_num_conv))
 | 
| 69597 |    125 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    126 |       transitive'
 | 
|  |    127 |         (inst [] [m, n] @{thm add_num_simps(8) [meta]})
 | 
|  |    128 |         (cong1 (args2 add_num_conv))
 | 
| 69597 |    129 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    130 |       transitive'
 | 
|  |    131 |         (inst [] [m, n] @{thm add_num_simps(9) [meta]})
 | 
|  |    132 |         (cong1 (cong2' add_num_conv (args2 add_num_conv) Thm.reflexive)));
 | 
| 64998 |    133 | \<close>
 | 
| 64962 |    134 | 
 | 
| 64998 |    135 | ML \<open>
 | 
| 64962 |    136 | fun BitM_conv m = (case strip_app m of
 | 
| 69597 |    137 |     (\<^const_name>\<open>Num.One\<close>, []) => @{thm BitM.simps(1) [meta]}
 | 
|  |    138 |   | (\<^const_name>\<open>Num.Bit0\<close>, [n]) =>
 | 
| 64962 |    139 |       transitive'
 | 
|  |    140 |         (inst [] [n] @{thm BitM.simps(2) [meta]})
 | 
|  |    141 |         (cong1 (args1 BitM_conv))
 | 
| 69597 |    142 |   | (\<^const_name>\<open>Num.Bit1\<close>, [n]) =>
 | 
| 64962 |    143 |       inst [] [n] @{thm BitM.simps(3) [meta]});
 | 
| 64998 |    144 | \<close>
 | 
| 64962 |    145 | 
 | 
|  |    146 | lemma dbl_neg_numeral:
 | 
|  |    147 |   "Num.dbl (- Num.numeral k) = - Num.numeral (Num.Bit0 k)"
 | 
|  |    148 |   by simp
 | 
|  |    149 | 
 | 
| 64998 |    150 | ML \<open>
 | 
| 64962 |    151 | fun dbl_conv a =
 | 
|  |    152 |   let
 | 
|  |    153 |     val dbl_neg_numeral_a = inst [a] [] @{thm dbl_neg_numeral [meta]};
 | 
|  |    154 |     val dbl_0_a = inst [a] [] @{thm dbl_simps(2) [meta]};
 | 
|  |    155 |     val dbl_numeral_a = inst [a] [] @{thm dbl_simps(5) [meta]}
 | 
|  |    156 |   in
 | 
|  |    157 |     fn n =>
 | 
|  |    158 |       case strip_numeral n of
 | 
| 69597 |    159 |         (\<^const_name>\<open>zero_class.zero\<close>, []) => dbl_0_a
 | 
|  |    160 |       | (\<^const_name>\<open>numeral\<close>, [k]) => inst [] [k] dbl_numeral_a
 | 
|  |    161 |       | (\<^const_name>\<open>uminus\<close>, [k]) => inst [] [k] dbl_neg_numeral_a
 | 
| 64962 |    162 |   end;
 | 
| 64998 |    163 | \<close>
 | 
| 64962 |    164 | 
 | 
|  |    165 | lemma dbl_inc_neg_numeral:
 | 
|  |    166 |   "Num.dbl_inc (- Num.numeral k) = - Num.numeral (Num.BitM k)"
 | 
|  |    167 |   by simp
 | 
|  |    168 | 
 | 
| 64998 |    169 | ML \<open>
 | 
| 64962 |    170 | fun dbl_inc_conv a =
 | 
|  |    171 |   let
 | 
|  |    172 |     val dbl_inc_neg_numeral_a = inst [a] [] @{thm dbl_inc_neg_numeral [meta]};
 | 
|  |    173 |     val dbl_inc_0_a = inst [a] [] @{thm dbl_inc_simps(2) [folded numeral_One, meta]};
 | 
|  |    174 |     val dbl_inc_numeral_a = inst [a] [] @{thm dbl_inc_simps(5) [meta]};
 | 
|  |    175 |   in
 | 
|  |    176 |     fn n =>
 | 
|  |    177 |       case strip_numeral n of
 | 
| 69597 |    178 |         (\<^const_name>\<open>zero_class.zero\<close>, []) => dbl_inc_0_a
 | 
|  |    179 |       | (\<^const_name>\<open>numeral\<close>, [k]) => inst [] [k] dbl_inc_numeral_a
 | 
|  |    180 |       | (\<^const_name>\<open>uminus\<close>, [k]) =>
 | 
| 64962 |    181 |           transitive'
 | 
|  |    182 |             (inst [] [k] dbl_inc_neg_numeral_a)
 | 
|  |    183 |             (cong1 (cong1 (args1 BitM_conv)))
 | 
|  |    184 |   end;
 | 
| 64998 |    185 | \<close>
 | 
| 64962 |    186 | 
 | 
|  |    187 | lemma dbl_dec_neg_numeral:
 | 
|  |    188 |   "Num.dbl_dec (- Num.numeral k) = - Num.numeral (Num.Bit1 k)"
 | 
|  |    189 |   by simp
 | 
|  |    190 | 
 | 
| 64998 |    191 | ML \<open>
 | 
| 64962 |    192 | fun dbl_dec_conv a =
 | 
|  |    193 |   let
 | 
|  |    194 |     val dbl_dec_neg_numeral_a = inst [a] [] @{thm dbl_dec_neg_numeral [meta]};
 | 
|  |    195 |     val dbl_dec_0_a = inst [a] [] @{thm dbl_dec_simps(2) [folded numeral_One, meta]};
 | 
|  |    196 |     val dbl_dec_numeral_a = inst [a] [] @{thm dbl_dec_simps(5) [meta]};
 | 
|  |    197 |   in
 | 
|  |    198 |     fn n =>
 | 
|  |    199 |       case strip_numeral n of
 | 
| 69597 |    200 |         (\<^const_name>\<open>zero_class.zero\<close>, []) => dbl_dec_0_a
 | 
|  |    201 |       | (\<^const_name>\<open>uminus\<close>, [k]) => inst [] [k] dbl_dec_neg_numeral_a
 | 
|  |    202 |       | (\<^const_name>\<open>numeral\<close>, [k]) =>
 | 
| 64962 |    203 |           transitive'
 | 
|  |    204 |             (inst [] [k] dbl_dec_numeral_a)
 | 
|  |    205 |             (cong1 (args1 BitM_conv))
 | 
|  |    206 |   end;
 | 
| 64998 |    207 | \<close>
 | 
| 64962 |    208 | 
 | 
| 64998 |    209 | ML \<open>
 | 
| 64962 |    210 | fun sub_conv a =
 | 
|  |    211 |   let
 | 
|  |    212 |     val [sub_One_One, sub_One_Bit0, sub_One_Bit1,
 | 
|  |    213 |          sub_Bit0_One, sub_Bit1_One, sub_Bit0_Bit0,
 | 
|  |    214 |          sub_Bit0_Bit1, sub_Bit1_Bit0, sub_Bit1_Bit1] =
 | 
|  |    215 |       map (inst [a] []) @{thms sub_num_simps [meta]};
 | 
|  |    216 |     val dbl_conv_a = dbl_conv a;
 | 
|  |    217 |     val dbl_inc_conv_a = dbl_inc_conv a;
 | 
|  |    218 |     val dbl_dec_conv_a = dbl_dec_conv a;
 | 
|  |    219 | 
 | 
|  |    220 |     fun conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    221 |         ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    222 |           sub_One_One
 | 
| 69597 |    223 |       | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit0\<close>, [l])) =>
 | 
| 64962 |    224 |           transitive'
 | 
|  |    225 |             (inst [] [l] sub_One_Bit0)
 | 
|  |    226 |             (cong1 (cong1 (args1 BitM_conv)))
 | 
| 69597 |    227 |       | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit1\<close>, [l])) =>
 | 
| 64962 |    228 |           inst [] [l] sub_One_Bit1
 | 
| 69597 |    229 |       | ((\<^const_name>\<open>Num.Bit0\<close>, [k]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    230 |           transitive'
 | 
|  |    231 |             (inst [] [k] sub_Bit0_One)
 | 
|  |    232 |             (cong1 (args1 BitM_conv))
 | 
| 69597 |    233 |       | ((\<^const_name>\<open>Num.Bit1\<close>, [k]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    234 |           inst [] [k] sub_Bit1_One
 | 
| 69597 |    235 |       | ((\<^const_name>\<open>Num.Bit0\<close>, [k]), (\<^const_name>\<open>Num.Bit0\<close>, [l])) =>
 | 
| 64962 |    236 |           transitive'
 | 
|  |    237 |             (inst [] [k, l] sub_Bit0_Bit0)
 | 
|  |    238 |             (cong1' dbl_conv_a (args2 conv))
 | 
| 69597 |    239 |       | ((\<^const_name>\<open>Num.Bit0\<close>, [k]), (\<^const_name>\<open>Num.Bit1\<close>, [l])) =>
 | 
| 64962 |    240 |           transitive'
 | 
|  |    241 |             (inst [] [k, l] sub_Bit0_Bit1)
 | 
|  |    242 |             (cong1' dbl_dec_conv_a (args2 conv))
 | 
| 69597 |    243 |       | ((\<^const_name>\<open>Num.Bit1\<close>, [k]), (\<^const_name>\<open>Num.Bit0\<close>, [l])) =>
 | 
| 64962 |    244 |           transitive'
 | 
|  |    245 |             (inst [] [k, l] sub_Bit1_Bit0)
 | 
|  |    246 |             (cong1' dbl_inc_conv_a (args2 conv))
 | 
| 69597 |    247 |       | ((\<^const_name>\<open>Num.Bit1\<close>, [k]), (\<^const_name>\<open>Num.Bit1\<close>, [l])) =>
 | 
| 64962 |    248 |           transitive'
 | 
|  |    249 |             (inst [] [k, l] sub_Bit1_Bit1)
 | 
|  |    250 |             (cong1' dbl_conv_a (args2 conv)))
 | 
|  |    251 |   in conv end;
 | 
| 64998 |    252 | \<close>
 | 
| 64962 |    253 | 
 | 
| 64998 |    254 | ML \<open>
 | 
| 64962 |    255 | fun expand1 a =
 | 
|  |    256 |   let val numeral_1_eq_1_a = inst [a] [] @{thm numeral_One [meta, symmetric]}
 | 
|  |    257 |   in
 | 
|  |    258 |     fn n =>
 | 
|  |    259 |       case Thm.term_of n of
 | 
| 74397 |    260 |         \<^Const_>\<open>one_class.one _\<close> => numeral_1_eq_1_a
 | 
|  |    261 |       | \<^Const_>\<open>uminus _ for \<^Const_>\<open>one_class.one _\<close>\<close> =>
 | 
| 64962 |    262 |           Thm.combination (Thm.reflexive (Thm.dest_fun n)) numeral_1_eq_1_a
 | 
| 74397 |    263 |       | \<^Const_>\<open>zero_class.zero _\<close> => Thm.reflexive n
 | 
|  |    264 |       | \<^Const_>\<open>numeral _ for _\<close> => Thm.reflexive n
 | 
|  |    265 |       | \<^Const_>\<open>uminus _ for \<^Const_>\<open>numeral _ for _\<close>\<close> => Thm.reflexive n
 | 
| 64962 |    266 |       | _ => err "expand1" n
 | 
|  |    267 |   end;
 | 
|  |    268 | 
 | 
|  |    269 | fun norm1_eq a =
 | 
|  |    270 |   let val numeral_1_eq_1_a = inst [a] [] @{thm numeral_One [meta]}
 | 
|  |    271 |   in
 | 
|  |    272 |     fn eq =>
 | 
|  |    273 |       case Thm.term_of (Thm.rhs_of eq) of
 | 
| 74397 |    274 |         \<^Const_>\<open>Num.numeral _ for \<^Const_>\<open>Num.One\<close>\<close> => Thm.transitive eq numeral_1_eq_1_a
 | 
|  |    275 |       | \<^Const_>\<open>uminus _ for \<^Const_>\<open>Num.numeral _ for \<^Const_>\<open>Num.One\<close>\<close>\<close> =>
 | 
| 64962 |    276 |             Thm.transitive eq
 | 
|  |    277 |               (Thm.combination (Thm.reflexive (Thm.dest_fun (Thm.rhs_of eq)))
 | 
|  |    278 |                  numeral_1_eq_1_a)
 | 
|  |    279 |       | _ => eq
 | 
|  |    280 |   end;
 | 
| 64998 |    281 | \<close>
 | 
| 64962 |    282 | 
 | 
| 64998 |    283 | ML \<open>
 | 
| 64962 |    284 | fun plus_conv f a =
 | 
|  |    285 |   let
 | 
|  |    286 |     val add_0_a = inst [a] [] @{thm add_0 [meta]};
 | 
|  |    287 |     val add_0_right_a = inst [a] [] @{thm add_0_right [meta]};
 | 
|  |    288 |     val numeral_plus_numeral_a = inst [a] [] @{thm numeral_plus_numeral [meta]};
 | 
|  |    289 |     val expand1_a = expand1 a;
 | 
|  |    290 | 
 | 
|  |    291 |     fun conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    292 |         ((\<^const_name>\<open>zero_class.zero\<close>, []), _) => inst [] [n] add_0_a
 | 
|  |    293 |       | (_, (\<^const_name>\<open>zero_class.zero\<close>, [])) => inst [] [m] add_0_right_a
 | 
|  |    294 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    295 |           transitive'
 | 
|  |    296 |             (inst [] [m, n] numeral_plus_numeral_a)
 | 
|  |    297 |             (cong1 (args2 add_num_conv))
 | 
|  |    298 |       | _ => cong2'' (f conv) (expand1_a m) (expand1_a n))
 | 
|  |    299 |   in f conv end;
 | 
|  |    300 | 
 | 
| 69597 |    301 | val nat_plus_conv = plus_conv I \<^ctyp>\<open>nat\<close>;
 | 
| 64998 |    302 | \<close>
 | 
| 64962 |    303 | 
 | 
|  |    304 | lemma neg_numeral_plus_neg_numeral:
 | 
|  |    305 |   "- Num.numeral m + - Num.numeral n = (- Num.numeral (m + n) ::'a::neg_numeral)"
 | 
|  |    306 |   by simp
 | 
|  |    307 | 
 | 
| 64998 |    308 | ML \<open>
 | 
| 64962 |    309 | fun plus_neg_conv a =
 | 
|  |    310 |   let
 | 
|  |    311 |     val numeral_plus_neg_numeral_a =
 | 
|  |    312 |       inst [a] [] @{thm add_neg_numeral_simps(1) [meta]};
 | 
|  |    313 |     val neg_numeral_plus_numeral_a =
 | 
|  |    314 |       inst [a] [] @{thm add_neg_numeral_simps(2) [meta]};
 | 
|  |    315 |     val neg_numeral_plus_neg_numeral_a =
 | 
|  |    316 |       inst [a] [] @{thm neg_numeral_plus_neg_numeral [meta]};
 | 
|  |    317 |     val sub_conv_a = sub_conv a;
 | 
|  |    318 |   in
 | 
|  |    319 |     fn conv => fn m => fn n => 
 | 
|  |    320 |       case (strip_numeral m, strip_numeral n) of
 | 
| 69597 |    321 |         ((\<^const_name>\<open>Num.numeral\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    322 |           Thm.transitive
 | 
|  |    323 |             (inst [] [m, n] numeral_plus_neg_numeral_a)
 | 
|  |    324 |             (sub_conv_a m n)
 | 
| 69597 |    325 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>Num.numeral\<close>, [n])) =>
 | 
| 64962 |    326 |           Thm.transitive
 | 
|  |    327 |             (inst [] [m, n] neg_numeral_plus_numeral_a)
 | 
|  |    328 |             (sub_conv_a n m)
 | 
| 69597 |    329 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    330 |           transitive'
 | 
|  |    331 |             (inst [] [m, n] neg_numeral_plus_neg_numeral_a)
 | 
|  |    332 |             (cong1 (cong1 (args2 add_num_conv)))
 | 
|  |    333 |       | _ => conv m n
 | 
|  |    334 |   end;
 | 
|  |    335 | 
 | 
|  |    336 | fun plus_conv' a = norm1_eq a oo plus_conv (plus_neg_conv a) a;
 | 
|  |    337 | 
 | 
| 69597 |    338 | val int_plus_conv = plus_conv' \<^ctyp>\<open>int\<close>;
 | 
| 64998 |    339 | \<close>
 | 
| 64962 |    340 | 
 | 
|  |    341 | lemma minus_one: "- 1 = - 1" by simp
 | 
|  |    342 | lemma minus_numeral: "- numeral b = - numeral b" by simp
 | 
|  |    343 | 
 | 
| 64998 |    344 | ML \<open>
 | 
| 64962 |    345 | fun uminus_conv a =
 | 
|  |    346 |   let
 | 
|  |    347 |     val minus_zero_a = inst [a] [] @{thm minus_zero [meta]};
 | 
|  |    348 |     val minus_one_a = inst [a] [] @{thm minus_one [meta]};
 | 
|  |    349 |     val minus_numeral_a = inst [a] [] @{thm minus_numeral [meta]};
 | 
|  |    350 |     val minus_minus_a = inst [a] [] @{thm minus_minus [meta]}
 | 
|  |    351 |   in
 | 
|  |    352 |     fn n =>
 | 
|  |    353 |       case strip_app n of
 | 
| 69597 |    354 |         (\<^const_name>\<open>zero_class.zero\<close>, []) => minus_zero_a
 | 
|  |    355 |       | (\<^const_name>\<open>one_class.one\<close>, []) => minus_one_a
 | 
|  |    356 |       | (\<^const_name>\<open>Num.numeral\<close>, [m]) => inst [] [m] minus_numeral_a
 | 
|  |    357 |       | (\<^const_name>\<open>uminus\<close>, [m]) => inst [] [m] minus_minus_a
 | 
| 64962 |    358 |   end;
 | 
|  |    359 | 
 | 
| 69597 |    360 | val int_neg_conv = uminus_conv \<^ctyp>\<open>int\<close>;
 | 
| 64998 |    361 | \<close>
 | 
| 64962 |    362 | 
 | 
| 64998 |    363 | ML \<open>
 | 
| 64962 |    364 | fun minus_conv a =
 | 
|  |    365 |   let
 | 
|  |    366 |     val [numeral_minus_numeral_a, numeral_minus_neg_numeral_a,
 | 
|  |    367 |          neg_numeral_minus_numeral_a, neg_numeral_minus_neg_numeral_a] =
 | 
|  |    368 |       map (inst [a] []) @{thms diff_numeral_simps [meta]};
 | 
|  |    369 |     val diff_0_a = inst [a] [] @{thm diff_0 [meta]};
 | 
|  |    370 |     val diff_0_right_a = inst [a] [] @{thm diff_0_right [meta]};
 | 
|  |    371 |     val sub_conv_a = sub_conv a;
 | 
|  |    372 |     val uminus_conv_a = uminus_conv a;
 | 
|  |    373 |     val expand1_a = expand1 a;
 | 
|  |    374 |     val norm1_eq_a = norm1_eq a;
 | 
|  |    375 | 
 | 
|  |    376 |     fun conv m n = (case (strip_numeral m, strip_numeral n) of
 | 
| 69597 |    377 |         ((\<^const_name>\<open>zero_class.zero\<close>, []), _) =>
 | 
| 64962 |    378 |           Thm.transitive (inst [] [n] diff_0_a) (uminus_conv_a n)
 | 
| 69597 |    379 |       | (_, (\<^const_name>\<open>zero_class.zero\<close>, [])) => inst [] [m] diff_0_right_a
 | 
|  |    380 |       | ((\<^const_name>\<open>Num.numeral\<close>, [m]), (\<^const_name>\<open>Num.numeral\<close>, [n])) =>
 | 
| 64962 |    381 |           Thm.transitive
 | 
|  |    382 |             (inst [] [m, n] numeral_minus_numeral_a)
 | 
|  |    383 |             (sub_conv_a m n)
 | 
| 69597 |    384 |       | ((\<^const_name>\<open>Num.numeral\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    385 |           transitive'
 | 
|  |    386 |             (inst [] [m, n] numeral_minus_neg_numeral_a)
 | 
|  |    387 |             (cong1 (args2 add_num_conv))
 | 
| 69597 |    388 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>Num.numeral\<close>, [n])) =>
 | 
| 64962 |    389 |           transitive'
 | 
|  |    390 |             (inst [] [m, n] neg_numeral_minus_numeral_a)
 | 
|  |    391 |             (cong1 (cong1 (args2 add_num_conv)))
 | 
| 69597 |    392 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    393 |           Thm.transitive
 | 
|  |    394 |             (inst [] [m, n] neg_numeral_minus_neg_numeral_a)
 | 
|  |    395 |             (sub_conv_a n m)
 | 
|  |    396 |       | _ => cong2'' conv (expand1_a m) (expand1_a n))
 | 
|  |    397 |   in norm1_eq_a oo conv end;
 | 
|  |    398 | 
 | 
| 69597 |    399 | val int_minus_conv = minus_conv \<^ctyp>\<open>int\<close>;
 | 
| 64998 |    400 | \<close>
 | 
| 64962 |    401 | 
 | 
| 64998 |    402 | ML \<open>
 | 
| 69597 |    403 | val int_numeral = Thm.apply \<^cterm>\<open>numeral :: num \<Rightarrow> int\<close>;
 | 
| 64962 |    404 | 
 | 
| 69597 |    405 | val nat_minus_refl = Thm.reflexive \<^cterm>\<open>minus :: nat \<Rightarrow> nat \<Rightarrow> nat\<close>;
 | 
| 64962 |    406 | 
 | 
| 69597 |    407 | val expand1_nat = expand1 \<^ctyp>\<open>nat\<close>;
 | 
| 64962 |    408 | 
 | 
|  |    409 | fun nat_minus_conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    410 |     ((\<^const_name>\<open>zero_class.zero\<close>, []), _) =>
 | 
| 64962 |    411 |       inst [] [n] @{thm diff_0_eq_0 [meta]}
 | 
| 69597 |    412 |   | (_, (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    413 |       inst [] [m] @{thm minus_nat.diff_0 [meta]}
 | 
| 69597 |    414 |   | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    415 |       transitive'
 | 
|  |    416 |         (inst [] [m, n] @{thm diff_nat_numeral [meta]})
 | 
|  |    417 |         (cong1' nat_conv (args2 int_minus_conv))
 | 
|  |    418 |   | _ => cong2'' nat_minus_conv (expand1_nat m) (expand1_nat n));
 | 
| 64998 |    419 | \<close>
 | 
| 64962 |    420 | 
 | 
| 64998 |    421 | ML \<open>
 | 
| 64962 |    422 | fun mult_num_conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    423 |     (_, (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    424 |       inst [] [m] @{thm mult_num_simps(1) [meta]}
 | 
| 69597 |    425 |   | ((\<^const_name>\<open>Num.One\<close>, []), _) =>
 | 
| 64962 |    426 |       inst [] [n] @{thm mult_num_simps(2) [meta]}
 | 
| 69597 |    427 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    428 |       transitive'
 | 
|  |    429 |         (inst [] [m, n] @{thm mult_num_simps(3) [meta]})
 | 
|  |    430 |         (cong1 (cong1 (args2 mult_num_conv)))
 | 
| 69597 |    431 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n'])) =>
 | 
| 64962 |    432 |       transitive'
 | 
|  |    433 |         (inst [] [m, n'] @{thm mult_num_simps(4) [meta]})
 | 
|  |    434 |         (cong1 (args2 mult_num_conv))
 | 
| 69597 |    435 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m']), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    436 |       transitive'
 | 
|  |    437 |         (inst [] [m', n] @{thm mult_num_simps(5) [meta]})
 | 
|  |    438 |         (cong1 (args2 mult_num_conv))
 | 
| 69597 |    439 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    440 |       transitive'
 | 
|  |    441 |         (inst [] [m, n] @{thm mult_num_simps(6) [meta]})
 | 
|  |    442 |         (cong1 (cong2' add_num_conv
 | 
|  |    443 |            (args2 add_num_conv)
 | 
|  |    444 |            (cong1 (args2 mult_num_conv)))));
 | 
| 64998 |    445 | \<close>
 | 
| 64962 |    446 | 
 | 
| 64998 |    447 | ML \<open>
 | 
| 64962 |    448 | fun mult_conv f a =
 | 
|  |    449 |   let
 | 
|  |    450 |     val mult_zero_left_a = inst [a] [] @{thm mult_zero_left [meta]};
 | 
|  |    451 |     val mult_zero_right_a = inst [a] [] @{thm mult_zero_right [meta]};
 | 
|  |    452 |     val numeral_times_numeral_a = inst [a] [] @{thm numeral_times_numeral [meta]};
 | 
|  |    453 |     val expand1_a = expand1 a;
 | 
|  |    454 |     val norm1_eq_a = norm1_eq a;
 | 
|  |    455 | 
 | 
|  |    456 |     fun conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    457 |         ((\<^const_name>\<open>zero_class.zero\<close>, []), _) => inst [] [n] mult_zero_left_a
 | 
|  |    458 |       | (_, (\<^const_name>\<open>zero_class.zero\<close>, [])) => inst [] [m] mult_zero_right_a
 | 
|  |    459 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    460 |           transitive'
 | 
|  |    461 |             (inst [] [m, n] numeral_times_numeral_a)
 | 
|  |    462 |             (cong1 (args2 mult_num_conv))
 | 
|  |    463 |       | _ => cong2'' (f conv) (expand1_a m) (expand1_a n))
 | 
|  |    464 |   in norm1_eq_a oo f conv end;
 | 
|  |    465 | 
 | 
| 69597 |    466 | val nat_mult_conv = mult_conv I \<^ctyp>\<open>nat\<close>;
 | 
| 64998 |    467 | \<close>
 | 
| 64962 |    468 | 
 | 
| 64998 |    469 | ML \<open>
 | 
| 64962 |    470 | fun mult_neg_conv a =
 | 
|  |    471 |   let
 | 
|  |    472 |     val [neg_numeral_times_neg_numeral_a, neg_numeral_times_numeral_a,
 | 
|  |    473 |          numeral_times_neg_numeral_a] =
 | 
|  |    474 |       map (inst [a] []) @{thms mult_neg_numeral_simps [meta]};
 | 
|  |    475 |   in
 | 
|  |    476 |     fn conv => fn m => fn n =>
 | 
|  |    477 |       case (strip_numeral m, strip_numeral n) of
 | 
| 69597 |    478 |         ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    479 |           transitive'
 | 
|  |    480 |             (inst [] [m, n] neg_numeral_times_neg_numeral_a)
 | 
|  |    481 |             (cong1 (args2 mult_num_conv))
 | 
| 69597 |    482 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    483 |           transitive'
 | 
|  |    484 |             (inst [] [m, n] neg_numeral_times_numeral_a)
 | 
|  |    485 |             (cong1 (cong1 (args2 mult_num_conv)))
 | 
| 69597 |    486 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    487 |           transitive'
 | 
|  |    488 |             (inst [] [m, n] numeral_times_neg_numeral_a)
 | 
|  |    489 |             (cong1 (cong1 (args2 mult_num_conv)))
 | 
|  |    490 |       | _ => conv m n
 | 
|  |    491 |   end;
 | 
|  |    492 | 
 | 
|  |    493 | fun mult_conv' a = mult_conv (mult_neg_conv a) a;
 | 
|  |    494 | 
 | 
| 69597 |    495 | val int_mult_conv = mult_conv' \<^ctyp>\<open>int\<close>;
 | 
| 64998 |    496 | \<close>
 | 
| 64962 |    497 | 
 | 
| 64998 |    498 | ML \<open>
 | 
| 64962 |    499 | fun eq_num_conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    500 |     ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    501 |       @{thm eq_num_simps(1) [meta]}
 | 
| 69597 |    502 |   | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    503 |       inst [] [n] @{thm eq_num_simps(2) [meta]}
 | 
| 69597 |    504 |   | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    505 |       inst [] [n] @{thm eq_num_simps(3) [meta]}
 | 
| 69597 |    506 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    507 |       inst [] [m] @{thm eq_num_simps(4) [meta]}
 | 
| 69597 |    508 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    509 |       inst [] [m] @{thm eq_num_simps(5) [meta]}
 | 
| 69597 |    510 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    511 |       Thm.transitive
 | 
|  |    512 |         (inst [] [m, n] @{thm eq_num_simps(6) [meta]})
 | 
|  |    513 |         (eq_num_conv m n)
 | 
| 69597 |    514 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    515 |       inst [] [m, n] @{thm eq_num_simps(7) [meta]}
 | 
| 69597 |    516 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    517 |       inst [] [m, n] @{thm eq_num_simps(8) [meta]}
 | 
| 69597 |    518 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    519 |       Thm.transitive
 | 
|  |    520 |         (inst [] [m, n] @{thm eq_num_simps(9) [meta]})
 | 
|  |    521 |         (eq_num_conv m n));
 | 
| 64998 |    522 | \<close>
 | 
| 64962 |    523 | 
 | 
| 64998 |    524 | ML \<open>
 | 
| 64962 |    525 | fun eq_conv f a =
 | 
|  |    526 |   let
 | 
|  |    527 |     val zero_eq_zero_a = inst [a] [] @{thm refl [of 0, THEN Eq_TrueI]};
 | 
|  |    528 |     val zero_neq_numeral_a =
 | 
|  |    529 |       inst [a] [] @{thm zero_neq_numeral [THEN Eq_FalseI]};
 | 
|  |    530 |     val numeral_neq_zero_a =
 | 
|  |    531 |       inst [a] [] @{thm numeral_neq_zero [THEN Eq_FalseI]};
 | 
|  |    532 |     val numeral_eq_iff_a = inst [a] [] @{thm numeral_eq_iff [meta]};
 | 
|  |    533 |     val expand1_a = expand1 a;
 | 
|  |    534 | 
 | 
|  |    535 |     fun conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    536 |         ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    537 |           zero_eq_zero_a
 | 
| 69597 |    538 |       | ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    539 |           inst [] [n] zero_neq_numeral_a
 | 
| 69597 |    540 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    541 |           inst [] [m] numeral_neq_zero_a
 | 
| 69597 |    542 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    543 |           Thm.transitive
 | 
|  |    544 |             (inst [] [m, n] numeral_eq_iff_a)
 | 
|  |    545 |             (eq_num_conv m n)
 | 
|  |    546 |       | _ => cong2'' (f conv) (expand1_a m) (expand1_a n))
 | 
|  |    547 |   in f conv end;
 | 
|  |    548 | 
 | 
| 69597 |    549 | val nat_eq_conv = eq_conv I \<^ctyp>\<open>nat\<close>;
 | 
| 64998 |    550 | \<close>
 | 
| 64962 |    551 | 
 | 
| 64998 |    552 | ML \<open>
 | 
| 64962 |    553 | fun eq_neg_conv a =
 | 
|  |    554 |   let
 | 
|  |    555 |     val neg_numeral_neq_zero_a =
 | 
|  |    556 |       inst [a] [] @{thm neg_numeral_neq_zero [THEN Eq_FalseI]};
 | 
|  |    557 |     val zero_neq_neg_numeral_a =
 | 
|  |    558 |       inst [a] [] @{thm zero_neq_neg_numeral [THEN Eq_FalseI]};
 | 
|  |    559 |     val neg_numeral_neq_numeral_a =
 | 
|  |    560 |       inst [a] [] @{thm neg_numeral_neq_numeral [THEN Eq_FalseI]};
 | 
|  |    561 |     val numeral_neq_neg_numeral_a =
 | 
|  |    562 |       inst [a] [] @{thm numeral_neq_neg_numeral [THEN Eq_FalseI]};
 | 
|  |    563 |     val neg_numeral_eq_iff_a = inst [a] [] @{thm neg_numeral_eq_iff [meta]}
 | 
|  |    564 |   in
 | 
|  |    565 |     fn conv => fn m => fn n => 
 | 
|  |    566 |       case (strip_numeral m, strip_numeral n) of
 | 
| 69597 |    567 |         ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    568 |           inst [] [m] neg_numeral_neq_zero_a
 | 
| 69597 |    569 |       | ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    570 |           inst [] [n] zero_neq_neg_numeral_a
 | 
| 69597 |    571 |       | ((\<^const_name>\<open>Num.numeral\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    572 |           inst [] [m, n] numeral_neq_neg_numeral_a
 | 
| 69597 |    573 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>Num.numeral\<close>, [n])) =>
 | 
| 64962 |    574 |           inst [] [m, n] neg_numeral_neq_numeral_a
 | 
| 69597 |    575 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    576 |           Thm.transitive
 | 
|  |    577 |             (inst [] [m, n] neg_numeral_eq_iff_a)
 | 
|  |    578 |             (eq_num_conv m n)
 | 
|  |    579 |       | _ => conv m n
 | 
|  |    580 |   end;
 | 
|  |    581 | 
 | 
|  |    582 | fun eq_conv' a = eq_conv (eq_neg_conv a) a;
 | 
|  |    583 | 
 | 
| 69597 |    584 | val int_eq_conv = eq_conv' \<^ctyp>\<open>int\<close>;
 | 
| 64998 |    585 | \<close>
 | 
| 64962 |    586 | 
 | 
| 64998 |    587 | ML \<open>
 | 
| 64962 |    588 | fun le_num_conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    589 |     ((\<^const_name>\<open>Num.One\<close>, []), _) =>
 | 
| 64962 |    590 |       inst [] [n] @{thm le_num_simps(1) [meta]}
 | 
| 69597 |    591 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    592 |       inst [] [m] @{thm le_num_simps(2) [meta]}
 | 
| 69597 |    593 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    594 |       inst [] [m] @{thm le_num_simps(3) [meta]}
 | 
| 69597 |    595 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    596 |       Thm.transitive
 | 
|  |    597 |         (inst [] [m, n] @{thm le_num_simps(4) [meta]})
 | 
|  |    598 |         (le_num_conv m n)
 | 
| 69597 |    599 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    600 |       Thm.transitive
 | 
|  |    601 |         (inst [] [m, n] @{thm le_num_simps(5) [meta]})
 | 
|  |    602 |         (le_num_conv m n)
 | 
| 69597 |    603 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    604 |       Thm.transitive
 | 
|  |    605 |         (inst [] [m, n] @{thm le_num_simps(6) [meta]})
 | 
|  |    606 |         (le_num_conv m n)
 | 
| 69597 |    607 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    608 |       Thm.transitive
 | 
|  |    609 |         (inst [] [m, n] @{thm le_num_simps(7) [meta]})
 | 
|  |    610 |         (less_num_conv m n))
 | 
|  |    611 | 
 | 
|  |    612 | and less_num_conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    613 |     (_, (\<^const_name>\<open>Num.One\<close>, [])) =>
 | 
| 64962 |    614 |       inst [] [m] @{thm less_num_simps(1) [meta]}
 | 
| 69597 |    615 |   | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    616 |       inst [] [n] @{thm less_num_simps(2) [meta]}
 | 
| 69597 |    617 |   | ((\<^const_name>\<open>Num.One\<close>, []), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    618 |       inst [] [n] @{thm less_num_simps(3) [meta]}
 | 
| 69597 |    619 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    620 |       Thm.transitive
 | 
|  |    621 |         (inst [] [m, n] @{thm less_num_simps(4) [meta]})
 | 
|  |    622 |         (less_num_conv m n)
 | 
| 69597 |    623 |   | ((\<^const_name>\<open>Num.Bit0\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    624 |       Thm.transitive
 | 
|  |    625 |         (inst [] [m, n] @{thm less_num_simps(5) [meta]})
 | 
|  |    626 |         (le_num_conv m n)
 | 
| 69597 |    627 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit1\<close>, [n])) =>
 | 
| 64962 |    628 |       Thm.transitive
 | 
|  |    629 |         (inst [] [m, n] @{thm less_num_simps(6) [meta]})
 | 
|  |    630 |         (less_num_conv m n)
 | 
| 69597 |    631 |   | ((\<^const_name>\<open>Num.Bit1\<close>, [m]), (\<^const_name>\<open>Num.Bit0\<close>, [n])) =>
 | 
| 64962 |    632 |       Thm.transitive
 | 
|  |    633 |         (inst [] [m, n] @{thm less_num_simps(7) [meta]})
 | 
|  |    634 |         (less_num_conv m n));
 | 
| 64998 |    635 | \<close>
 | 
| 64962 |    636 | 
 | 
| 64998 |    637 | ML \<open>
 | 
| 64962 |    638 | fun le_conv f a =
 | 
|  |    639 |   let
 | 
|  |    640 |     val zero_le_zero_a = inst [a] [] @{thm order_refl [of 0, THEN Eq_TrueI]};
 | 
|  |    641 |     val zero_le_numeral_a =
 | 
|  |    642 |       inst [a] [] @{thm zero_le_numeral [THEN Eq_TrueI]};
 | 
|  |    643 |     val not_numeral_le_zero_a =
 | 
|  |    644 |       inst [a] [] @{thm not_numeral_le_zero [THEN Eq_FalseI]};
 | 
|  |    645 |     val numeral_le_iff_a = inst [a] [] @{thm numeral_le_iff [meta]};
 | 
|  |    646 |     val expand1_a = expand1 a;
 | 
|  |    647 | 
 | 
|  |    648 |     fun conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    649 |         ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    650 |           zero_le_zero_a
 | 
| 69597 |    651 |       | ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    652 |           inst [] [n] zero_le_numeral_a
 | 
| 69597 |    653 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    654 |           inst [] [m] not_numeral_le_zero_a
 | 
| 69597 |    655 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    656 |           Thm.transitive
 | 
|  |    657 |             (inst [] [m, n] numeral_le_iff_a)
 | 
|  |    658 |             (le_num_conv m n)
 | 
|  |    659 |       | _ => cong2'' (f conv) (expand1_a m) (expand1_a n))
 | 
|  |    660 |   in f conv end;
 | 
|  |    661 | 
 | 
| 69597 |    662 | val nat_le_conv = le_conv I \<^ctyp>\<open>nat\<close>;
 | 
| 64998 |    663 | \<close>
 | 
| 64962 |    664 | 
 | 
| 64998 |    665 | ML \<open>
 | 
| 64962 |    666 | fun le_neg_conv a =
 | 
|  |    667 |   let
 | 
|  |    668 |     val neg_numeral_le_zero_a =
 | 
|  |    669 |       inst [a] [] @{thm neg_numeral_le_zero [THEN Eq_TrueI]};
 | 
|  |    670 |     val not_zero_le_neg_numeral_a =
 | 
|  |    671 |       inst [a] [] @{thm not_zero_le_neg_numeral [THEN Eq_FalseI]};
 | 
|  |    672 |     val neg_numeral_le_numeral_a =
 | 
|  |    673 |       inst [a] [] @{thm neg_numeral_le_numeral [THEN Eq_TrueI]};
 | 
|  |    674 |     val not_numeral_le_neg_numeral_a =
 | 
|  |    675 |       inst [a] [] @{thm not_numeral_le_neg_numeral [THEN Eq_FalseI]};
 | 
|  |    676 |     val neg_numeral_le_iff_a = inst [a] [] @{thm neg_numeral_le_iff [meta]}
 | 
|  |    677 |   in
 | 
|  |    678 |     fn conv => fn m => fn n => 
 | 
|  |    679 |       case (strip_numeral m, strip_numeral n) of
 | 
| 69597 |    680 |         ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    681 |           inst [] [m] neg_numeral_le_zero_a
 | 
| 69597 |    682 |       | ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    683 |           inst [] [n] not_zero_le_neg_numeral_a
 | 
| 69597 |    684 |       | ((\<^const_name>\<open>Num.numeral\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    685 |           inst [] [m, n] not_numeral_le_neg_numeral_a
 | 
| 69597 |    686 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>Num.numeral\<close>, [n])) =>
 | 
| 64962 |    687 |           inst [] [m, n] neg_numeral_le_numeral_a
 | 
| 69597 |    688 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    689 |           Thm.transitive
 | 
|  |    690 |             (inst [] [m, n] neg_numeral_le_iff_a)
 | 
|  |    691 |             (le_num_conv n m)
 | 
|  |    692 |       | _ => conv m n
 | 
|  |    693 |   end;
 | 
|  |    694 | 
 | 
|  |    695 | fun le_conv' a = le_conv (le_neg_conv a) a;
 | 
|  |    696 | 
 | 
| 69597 |    697 | val int_le_conv = le_conv' \<^ctyp>\<open>int\<close>;
 | 
| 64998 |    698 | \<close>
 | 
| 64962 |    699 | 
 | 
| 64998 |    700 | ML \<open>
 | 
| 64962 |    701 | fun less_conv f a =
 | 
|  |    702 |   let
 | 
|  |    703 |     val not_zero_less_zero_a = inst [a] [] @{thm less_irrefl [of 0, THEN Eq_FalseI]};
 | 
|  |    704 |     val zero_less_numeral_a =
 | 
|  |    705 |       inst [a] [] @{thm zero_less_numeral [THEN Eq_TrueI]};
 | 
|  |    706 |     val not_numeral_less_zero_a =
 | 
|  |    707 |       inst [a] [] @{thm not_numeral_less_zero [THEN Eq_FalseI]};
 | 
|  |    708 |     val numeral_less_iff_a = inst [a] [] @{thm numeral_less_iff [meta]};
 | 
|  |    709 |     val expand1_a = expand1 a;
 | 
|  |    710 | 
 | 
|  |    711 |     fun conv m n = (case (strip_app m, strip_app n) of
 | 
| 69597 |    712 |         ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    713 |           not_zero_less_zero_a
 | 
| 69597 |    714 |       | ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    715 |           inst [] [n] zero_less_numeral_a
 | 
| 69597 |    716 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    717 |           inst [] [m] not_numeral_less_zero_a
 | 
| 69597 |    718 |       | ((\<^const_name>\<open>numeral\<close>, [m]), (\<^const_name>\<open>numeral\<close>, [n])) =>
 | 
| 64962 |    719 |           Thm.transitive
 | 
|  |    720 |             (inst [] [m, n] numeral_less_iff_a)
 | 
|  |    721 |             (less_num_conv m n)
 | 
|  |    722 |       | _ => cong2'' (f conv) (expand1_a m) (expand1_a n))
 | 
|  |    723 |   in f conv end;
 | 
|  |    724 | 
 | 
| 69597 |    725 | val nat_less_conv = less_conv I \<^ctyp>\<open>nat\<close>;
 | 
| 64998 |    726 | \<close>
 | 
| 64962 |    727 | 
 | 
| 64998 |    728 | ML \<open>
 | 
| 64962 |    729 | fun less_neg_conv a =
 | 
|  |    730 |   let
 | 
|  |    731 |     val neg_numeral_less_zero_a =
 | 
|  |    732 |       inst [a] [] @{thm neg_numeral_less_zero [THEN Eq_TrueI]};
 | 
|  |    733 |     val not_zero_less_neg_numeral_a =
 | 
|  |    734 |       inst [a] [] @{thm not_zero_less_neg_numeral [THEN Eq_FalseI]};
 | 
|  |    735 |     val neg_numeral_less_numeral_a =
 | 
|  |    736 |       inst [a] [] @{thm neg_numeral_less_numeral [THEN Eq_TrueI]};
 | 
|  |    737 |     val not_numeral_less_neg_numeral_a =
 | 
|  |    738 |       inst [a] [] @{thm not_numeral_less_neg_numeral [THEN Eq_FalseI]};
 | 
|  |    739 |     val neg_numeral_less_iff_a = inst [a] [] @{thm neg_numeral_less_iff [meta]}
 | 
|  |    740 |   in
 | 
|  |    741 |     fn conv => fn m => fn n => 
 | 
|  |    742 |       case (strip_numeral m, strip_numeral n) of
 | 
| 69597 |    743 |         ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>zero_class.zero\<close>, [])) =>
 | 
| 64962 |    744 |           inst [] [m] neg_numeral_less_zero_a
 | 
| 69597 |    745 |       | ((\<^const_name>\<open>zero_class.zero\<close>, []), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    746 |           inst [] [n] not_zero_less_neg_numeral_a
 | 
| 69597 |    747 |       | ((\<^const_name>\<open>Num.numeral\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    748 |           inst [] [m, n] not_numeral_less_neg_numeral_a
 | 
| 69597 |    749 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>Num.numeral\<close>, [n])) =>
 | 
| 64962 |    750 |           inst [] [m, n] neg_numeral_less_numeral_a
 | 
| 69597 |    751 |       | ((\<^const_name>\<open>uminus\<close>, [m]), (\<^const_name>\<open>uminus\<close>, [n])) =>
 | 
| 64962 |    752 |           Thm.transitive
 | 
|  |    753 |             (inst [] [m, n] neg_numeral_less_iff_a)
 | 
|  |    754 |             (less_num_conv n m)
 | 
|  |    755 |       | _ => conv m n
 | 
|  |    756 |   end;
 | 
|  |    757 | 
 | 
|  |    758 | fun less_conv' a = less_conv (less_neg_conv a) a;
 | 
|  |    759 | 
 | 
| 69597 |    760 | val int_less_conv = less_conv' \<^ctyp>\<open>int\<close>;
 | 
| 64998 |    761 | \<close>
 | 
| 64962 |    762 | 
 | 
| 64998 |    763 | ML \<open>
 | 
| 64962 |    764 | fun If_conv a =
 | 
|  |    765 |   let
 | 
|  |    766 |     val if_True = inst [a] [] @{thm if_True [meta]};
 | 
|  |    767 |     val if_False = inst [a] [] @{thm if_False [meta]}
 | 
|  |    768 |   in
 | 
|  |    769 |     fn p => fn x => fn y => fn ct =>
 | 
|  |    770 |       case strip_app ct of
 | 
| 69597 |    771 |         (\<^const_name>\<open>If\<close>, [cb, cx, cy]) =>
 | 
| 64962 |    772 |           let
 | 
|  |    773 |             val p_eq = p cb
 | 
|  |    774 |             val eq = Thm.combination (Thm.reflexive (Thm.dest_fun (Thm.dest_fun2 ct))) p_eq
 | 
|  |    775 |           in
 | 
|  |    776 |             case Thm.term_of (Thm.rhs_of p_eq) of
 | 
| 74397 |    777 |               \<^Const_>\<open>True\<close> =>
 | 
| 64962 |    778 |                 let
 | 
|  |    779 |                   val x_eq = x cx;
 | 
|  |    780 |                   val cx = Thm.rhs_of x_eq;
 | 
|  |    781 |                 in
 | 
|  |    782 |                   Thm.transitive
 | 
|  |    783 |                     (Thm.combination
 | 
|  |    784 |                        (Thm.combination eq x_eq)
 | 
|  |    785 |                        (Thm.reflexive cy))
 | 
|  |    786 |                     (inst [] [cx, cy] if_True)
 | 
|  |    787 |                 end
 | 
| 74397 |    788 |             | \<^Const_>\<open>False\<close> =>
 | 
| 64962 |    789 |                 let
 | 
|  |    790 |                   val y_eq = y cy;
 | 
|  |    791 |                   val cy = Thm.rhs_of y_eq;
 | 
|  |    792 |                 in
 | 
|  |    793 |                   Thm.transitive
 | 
|  |    794 |                     (Thm.combination
 | 
|  |    795 |                        (Thm.combination eq (Thm.reflexive cx))
 | 
|  |    796 |                        y_eq)
 | 
|  |    797 |                     (inst [] [cx, cy] if_False)
 | 
|  |    798 |                 end
 | 
|  |    799 |             | _ => err "If_conv" (Thm.rhs_of p_eq)
 | 
|  |    800 |           end
 | 
|  |    801 |   end;
 | 
| 64998 |    802 | \<close>
 | 
| 64962 |    803 | 
 | 
| 64998 |    804 | ML \<open>
 | 
| 64962 |    805 | fun drop_conv a =
 | 
|  |    806 |   let
 | 
|  |    807 |     val drop_0_a = inst [a] [] @{thm drop_0 [meta]};
 | 
|  |    808 |     val drop_Cons_a = inst [a] [] @{thm drop_Cons' [meta]};
 | 
|  |    809 |     val If_conv_a = If_conv (type_of_eqn drop_0_a);
 | 
|  |    810 | 
 | 
|  |    811 |     fun conv n ys = (case Thm.term_of n of
 | 
| 74397 |    812 |         \<^Const_>\<open>zero_class.zero _\<close> => inst [] [ys] drop_0_a
 | 
| 64962 |    813 |       | _ => (case strip_app ys of
 | 
| 69597 |    814 |           (\<^const_name>\<open>Cons\<close>, [x, xs]) =>
 | 
| 64962 |    815 |             transitive'
 | 
|  |    816 |               (inst [] [n, x, xs] drop_Cons_a)
 | 
|  |    817 |               (If_conv_a (args2 nat_eq_conv)
 | 
|  |    818 |                  Thm.reflexive
 | 
|  |    819 |                  (cong2' conv (args2 nat_minus_conv) Thm.reflexive))))
 | 
|  |    820 |   in conv end;
 | 
| 64998 |    821 | \<close>
 | 
| 64962 |    822 | 
 | 
| 64998 |    823 | ML \<open>
 | 
| 64962 |    824 | fun nth_conv a =
 | 
|  |    825 |   let
 | 
|  |    826 |     val nth_Cons_a = inst [a] [] @{thm nth_Cons' [meta]};
 | 
|  |    827 |     val If_conv_a = If_conv a;
 | 
|  |    828 | 
 | 
|  |    829 |     fun conv ys n = (case strip_app ys of
 | 
| 69597 |    830 |         (\<^const_name>\<open>Cons\<close>, [x, xs]) =>
 | 
| 64962 |    831 |           transitive'
 | 
|  |    832 |             (inst [] [x, xs, n] nth_Cons_a)
 | 
|  |    833 |             (If_conv_a (args2 nat_eq_conv)
 | 
|  |    834 |                Thm.reflexive
 | 
|  |    835 |                (cong2' conv Thm.reflexive (args2 nat_minus_conv))))
 | 
|  |    836 |   in conv end;
 | 
| 64998 |    837 | \<close>
 | 
| 64962 |    838 | 
 | 
|  |    839 | end
 |