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