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