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