src/Provers/simplifier.ML
author wenzelm
Wed, 01 Sep 1999 21:22:38 +0200
changeset 7423 28b48fcb0d55
parent 7273 d80b9be87535
child 7558 08b2d5c94b8a
permissions -rw-r--r--
Method.insert_tac; tuned comments;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1243
fa09705a5890 added global simpset
clasohm
parents: 983
diff changeset
     1
(*  Title:      Provers/simplifier.ML
1
c6a6e3db5353 changed addcongs to addeqcongs in simplifier.ML
nipkow
parents: 0
diff changeset
     2
    ID:         $Id$
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
     3
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
1
c6a6e3db5353 changed addcongs to addeqcongs in simplifier.ML
nipkow
parents: 0
diff changeset
     4
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
     5
Generic simplifier, suitable for most logics.  See Pure/thm.ML for the
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
     6
actual meta-level rewriting engine.
1
c6a6e3db5353 changed addcongs to addeqcongs in simplifier.ML
nipkow
parents: 0
diff changeset
     7
*)
1260
c76ac4a9dc2b renamed SS to Simpset; fixed bug in merge_ss
clasohm
parents: 1243
diff changeset
     8
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
     9
infix 4
4682
ff081854fc86 Added delspilts, Addsplits, Delsplits.
nipkow
parents: 4677
diff changeset
    10
  setsubgoaler setloop addloop delloop setSSolver addSSolver setSolver
4677
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
    11
  addSolver addsimps delsimps addeqcongs deleqcongs
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
    12
  setmksimps setmkeqTrue setmksym settermless addsimprocs delsimprocs;
2567
7a28e02e10b7 added addloop (and also documentation of addsolver
oheimb
parents: 2521
diff changeset
    13
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
    14
signature BASIC_SIMPLIFIER =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
sig
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
    16
  type simproc
3577
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    17
  val mk_simproc: string -> cterm list
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    18
    -> (Sign.sg -> thm list -> term -> thm option) -> simproc
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
  type simpset
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
    20
  val empty_ss: simpset
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
    21
  val rep_ss: simpset ->
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
    22
   {mss: meta_simpset,
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
    23
    subgoal_tac:        simpset -> int -> tactic,
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
    24
    loop_tacs:          (string * (int -> tactic))list,
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
    25
    unsafe_finish_tac: thm list -> int -> tactic,
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
    26
           finish_tac: thm list -> int -> tactic};
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
    27
  val print_ss: simpset -> unit
4366
4d84cdb0df42 added print_simpset;
wenzelm
parents: 4290
diff changeset
    28
  val print_simpset: theory -> unit
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    29
  val setsubgoaler: simpset *  (simpset -> int -> tactic) -> simpset
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    30
  val setloop:      simpset *             (int -> tactic) -> simpset
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
    31
  val addloop:      simpset *  (string * (int -> tactic)) -> simpset
4682
ff081854fc86 Added delspilts, Addsplits, Delsplits.
nipkow
parents: 4677
diff changeset
    32
  val delloop:      simpset *   string                    -> simpset
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    33
  val setSSolver:   simpset * (thm list -> int -> tactic) -> simpset
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    34
  val addSSolver:   simpset * (thm list -> int -> tactic) -> simpset
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    35
  val setSolver:    simpset * (thm list -> int -> tactic) -> simpset
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    36
  val addSolver:    simpset * (thm list -> int -> tactic) -> simpset
3577
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    37
  val setmksimps:   simpset * (thm -> thm list) -> simpset
4677
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
    38
  val setmkeqTrue:  simpset * (thm -> thm option) -> simpset
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
    39
  val setmksym:     simpset * (thm -> thm option) -> simpset
3577
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    40
  val settermless:  simpset * (term * term -> bool) -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    41
  val addsimps:     simpset * thm list -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    42
  val delsimps:     simpset * thm list -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    43
  val addeqcongs:   simpset * thm list -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    44
  val deleqcongs:   simpset * thm list -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    45
  val addsimprocs:  simpset * simproc list -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    46
  val delsimprocs:  simpset * simproc list -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    47
  val merge_ss:     simpset * simpset -> simpset
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
    48
  val prems_of_ss:  simpset -> thm list
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    49
  val simpset_ref_of_sg: Sign.sg -> simpset ref
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    50
  val simpset_ref_of: theory -> simpset ref
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    51
  val simpset_of_sg: Sign.sg -> simpset
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    52
  val simpset_of: theory -> simpset
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    53
  val SIMPSET: (simpset -> tactic) -> tactic
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    54
  val SIMPSET': (simpset -> 'a -> tactic) -> 'a -> tactic
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    55
  val simpset: unit -> simpset
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
    56
  val simpset_ref: unit -> simpset ref
1243
fa09705a5890 added global simpset
clasohm
parents: 983
diff changeset
    57
  val Addsimps: thm list -> unit
fa09705a5890 added global simpset
clasohm
parents: 983
diff changeset
    58
  val Delsimps: thm list -> unit
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
    59
  val Addsimprocs: simproc list -> unit
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
    60
  val Delsimprocs: simproc list -> unit
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    61
  val               simp_tac: simpset -> int -> tactic
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    62
  val           asm_simp_tac: simpset -> int -> tactic
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    63
  val          full_simp_tac: simpset -> int -> tactic
4722
d2e44673c21e The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents: 4713
diff changeset
    64
  val        asm_lr_simp_tac: simpset -> int -> tactic
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    65
  val      asm_full_simp_tac: simpset -> int -> tactic
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    66
  val safe_asm_full_simp_tac: simpset -> int -> tactic
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    67
  val               Simp_tac:            int -> tactic
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    68
  val           Asm_simp_tac:            int -> tactic
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    69
  val          Full_simp_tac:            int -> tactic
4722
d2e44673c21e The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents: 4713
diff changeset
    70
  val        Asm_lr_simp_tac:            int -> tactic
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
    71
  val      Asm_full_simp_tac:            int -> tactic
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
    72
  val          simplify: simpset -> thm -> thm
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
    73
  val      asm_simplify: simpset -> thm -> thm
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
    74
  val     full_simplify: simpset -> thm -> thm
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
    75
  val asm_full_simplify: simpset -> thm -> thm
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
    78
signature SIMPLIFIER =
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
    79
sig
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
    80
  include BASIC_SIMPLIFIER
5082
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
    81
  val          rewrite: simpset -> cterm -> thm
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
    82
  val      asm_rewrite: simpset -> cterm -> thm
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
    83
  val     full_rewrite: simpset -> cterm -> thm
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
    84
  val asm_full_rewrite: simpset -> cterm -> thm
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
    85
  val print_local_simpset: Proof.context -> unit
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
    86
  val get_local_simpset: Proof.context -> simpset
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
    87
  val put_local_simpset: simpset -> Proof.context -> Proof.context
4855
62bc389d6168 tuned setup;
wenzelm
parents: 4795
diff changeset
    88
  val simp_add_global: theory attribute
62bc389d6168 tuned setup;
wenzelm
parents: 4795
diff changeset
    89
  val simp_del_global: theory attribute
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
    90
  val simp_add_local: Proof.context attribute
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
    91
  val simp_del_local: Proof.context attribute
7177
6bb7ad30f3da change_simpset_of;
wenzelm
parents: 7132
diff changeset
    92
  val change_simpset_of: (simpset * 'a -> simpset) -> 'a -> theory -> theory
7273
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
    93
  val simp_modifiers: (Args.T list -> (Method.modifier * Args.T list)) list
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
    94
  val setup: (theory -> theory) list
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
    95
end;
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
    96
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
    97
structure Simplifier: SIMPLIFIER =
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
struct
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   100
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   101
(** simplification procedures **)
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   102
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   103
(* datatype simproc *)
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   104
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   105
datatype simproc =
3577
9715b6e3ec5f added prems argument to simplification procedures;
wenzelm
parents: 3557
diff changeset
   106
  Simproc of string * cterm list * (Sign.sg -> thm list -> term -> thm option) * stamp;
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   107
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   108
fun mk_simproc name lhss proc =
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   109
  Simproc (name, map (Thm.cterm_fun Logic.varify) lhss, proc, stamp ());
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   110
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   111
fun rep_simproc (Simproc args) = args;
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   112
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   113
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   114
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   115
(** simplification sets **)
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   116
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   117
(* type simpset *)
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   118
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
datatype simpset =
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   120
  Simpset of {
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   121
    mss: meta_simpset,
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
   122
    subgoal_tac:        simpset -> int -> tactic,
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   123
    loop_tacs:          (string * (int -> tactic))list,
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   124
    unsafe_finish_tac: thm list -> int -> tactic,
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   125
           finish_tac: thm list -> int -> tactic};
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   126
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   127
fun make_ss (mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac) =
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   128
  Simpset {mss = mss, subgoal_tac = subgoal_tac, loop_tacs = loop_tacs,
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   129
    unsafe_finish_tac = unsafe_finish_tac, finish_tac = finish_tac};
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
val empty_ss =
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   132
  let val mss = Thm.set_mk_sym (Thm.empty_mss, Some o symmetric_fun)
4677
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
   133
  in make_ss (mss, K (K no_tac), [], K (K no_tac), K (K no_tac)) end;
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   134
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   135
fun rep_ss (Simpset args) = args;
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   136
fun prems_of_ss (Simpset {mss, ...}) = Thm.prems_of_mss mss;
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   137
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   138
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   139
(* print simpsets *)
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   140
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   141
fun print_ss ss =
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   142
  let
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   143
    val Simpset {mss, ...} = ss;
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   144
    val {simps, procs, congs} = Thm.dest_mss mss;
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   145
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   146
    val pretty_thms = map Display.pretty_thm;
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   147
    fun pretty_proc (name, lhss) =
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   148
      Pretty.big_list (name ^ ":") (map Display.pretty_cterm lhss);
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   149
  in
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   150
    Pretty.writeln (Pretty.big_list "simplification rules:" (pretty_thms simps));
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   151
    Pretty.writeln (Pretty.big_list "simplification procedures:" (map pretty_proc procs));
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   152
    Pretty.writeln (Pretty.big_list "congruences:" (pretty_thms congs))
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   153
  end;
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   154
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   155
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   156
(* extend simpsets *)
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   157
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   158
fun (Simpset {mss, subgoal_tac = _, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   159
    setsubgoaler subgoal_tac =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   160
  make_ss (mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
   161
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   162
fun (Simpset {mss, subgoal_tac, loop_tacs = _, unsafe_finish_tac, finish_tac})
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   163
    setloop tac =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   164
  make_ss (mss, subgoal_tac, [("",tac)], unsafe_finish_tac, finish_tac);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   165
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   166
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
4741
7fcd106cb0ed addloop: added warning in case of overwriting a looper
oheimb
parents: 4739
diff changeset
   167
    addloop tac = make_ss (mss, subgoal_tac, 
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   168
        (case assoc_string (loop_tacs,(fst tac)) of None => () | Some x => 
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   169
         warning ("overwriting looper "^fst tac); overwrite(loop_tacs,tac)),
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   170
           unsafe_finish_tac, finish_tac);
2567
7a28e02e10b7 added addloop (and also documentation of addsolver
oheimb
parents: 2521
diff changeset
   171
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   172
fun (ss as Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   173
 delloop name =
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   174
  let val (del, rest) = partition (fn (n, _) => n = name) loop_tacs in
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   175
    if null del then (warning ("No such looper in simpset: " ^ name); ss)
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   176
    else make_ss (mss, subgoal_tac, rest, unsafe_finish_tac, finish_tac)
4682
ff081854fc86 Added delspilts, Addsplits, Delsplits.
nipkow
parents: 4677
diff changeset
   177
  end;
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   178
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, ...})
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   179
    setSSolver finish_tac =
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   180
  make_ss (mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
4682
ff081854fc86 Added delspilts, Addsplits, Delsplits.
nipkow
parents: 4677
diff changeset
   181
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   182
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   183
    addSSolver tac =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   184
  make_ss (mss, subgoal_tac, loop_tacs, unsafe_finish_tac,
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   185
	fn hyps => finish_tac hyps ORELSE' tac hyps);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   186
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   187
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac = _, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   188
    setSolver unsafe_finish_tac =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   189
  make_ss (mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   190
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   191
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   192
    addSolver tac =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   193
  make_ss (mss, subgoal_tac, loop_tacs, 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   194
    fn hyps => unsafe_finish_tac hyps ORELSE' tac hyps, finish_tac);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   195
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   196
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   197
    setmksimps mk_simps =
2645
9d3a3e62bf34 mk_rews: automatically includes strip_shyps, zero_var_indexes;
wenzelm
parents: 2629
diff changeset
   198
  make_ss (Thm.set_mk_rews (mss, map (Thm.strip_shyps o Drule.zero_var_indexes) o mk_simps),
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   199
    subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   200
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   201
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
4677
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
   202
    setmkeqTrue mk_eq_True =
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
   203
  make_ss (Thm.set_mk_eq_True (mss, mk_eq_True),
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   204
    subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
4677
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
   205
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   206
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
4677
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
   207
    setmksym mksym =
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
   208
  make_ss (Thm.set_mk_sym (mss, mksym),
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   209
    subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
4677
c4b07b8579fd Reorganized simplifier. May now reorient rules.
nipkow
parents: 4668
diff changeset
   210
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   211
fun (Simpset {mss, subgoal_tac, loop_tacs,  unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   212
    settermless termless =
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   213
  make_ss (Thm.set_termless (mss, termless), subgoal_tac, loop_tacs,
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   214
    unsafe_finish_tac, finish_tac);
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   215
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   216
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   217
    addsimps rews =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   218
  make_ss (Thm.add_simps (mss, rews), subgoal_tac, loop_tacs, 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   219
	   unsafe_finish_tac, finish_tac);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   220
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   221
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   222
    delsimps rews =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   223
  make_ss (Thm.del_simps (mss, rews), subgoal_tac, loop_tacs, 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   224
	   unsafe_finish_tac, finish_tac);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   225
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   226
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   227
    addeqcongs newcongs =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   228
  make_ss (Thm.add_congs (mss, newcongs), subgoal_tac, loop_tacs, 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   229
	   unsafe_finish_tac, finish_tac);
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   230
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   231
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   232
    deleqcongs oldcongs =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   233
  make_ss (Thm.del_congs (mss, oldcongs), subgoal_tac, loop_tacs, 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   234
	   unsafe_finish_tac, finish_tac);
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
   235
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   236
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   237
    addsimprocs simprocs =
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   238
  make_ss
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   239
    (Thm.add_simprocs (mss, map rep_simproc simprocs),
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   240
      subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   241
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   242
fun (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac})
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   243
    delsimprocs simprocs =
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   244
  make_ss
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   245
    (Thm.del_simprocs (mss, map rep_simproc simprocs),
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   246
      subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   247
7273
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   248
fun clear_ss (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac}) =
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   249
  make_ss (Thm.clear_mss mss, subgoal_tac, loop_tacs, unsafe_finish_tac, finish_tac);
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   250
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   251
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   252
(* merge simpsets *)    (*NOTE: ignores tactics of 2nd simpset (except loopers)*)
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   253
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   254
fun merge_ss
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   255
   (Simpset {mss = mss1, loop_tacs = loop_tacs1, subgoal_tac, unsafe_finish_tac, finish_tac},
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   256
    Simpset {mss = mss2, loop_tacs = loop_tacs2, ...}) =
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   257
  make_ss (Thm.merge_mss (mss1, mss2), subgoal_tac,
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   258
    merge_alists loop_tacs1 loop_tacs2, unsafe_finish_tac, finish_tac);
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   259
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   260
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   261
5024
d42ce3452dea adapted to TheoryDataFun interface;
wenzelm
parents: 5001
diff changeset
   262
(** global and local simpset data **)
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   263
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   264
(* theory data kind 'Provers/simpset' *)
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   265
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   266
structure GlobalSimpsetArgs =
5024
d42ce3452dea adapted to TheoryDataFun interface;
wenzelm
parents: 5001
diff changeset
   267
struct
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   268
  val name = "Provers/simpset";
5024
d42ce3452dea adapted to TheoryDataFun interface;
wenzelm
parents: 5001
diff changeset
   269
  type T = simpset ref;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
5024
d42ce3452dea adapted to TheoryDataFun interface;
wenzelm
parents: 5001
diff changeset
   271
  val empty = ref empty_ss;
6556
daa00919502b theory data: copy;
wenzelm
parents: 6534
diff changeset
   272
  fun copy (ref ss) = (ref ss): T;            (*create new reference!*)
daa00919502b theory data: copy;
wenzelm
parents: 6534
diff changeset
   273
  val prep_ext = copy;
5024
d42ce3452dea adapted to TheoryDataFun interface;
wenzelm
parents: 5001
diff changeset
   274
  fun merge (ref ss1, ref ss2) = ref (merge_ss (ss1, ss2));
d42ce3452dea adapted to TheoryDataFun interface;
wenzelm
parents: 5001
diff changeset
   275
  fun print _ (ref ss) = print_ss ss;
d42ce3452dea adapted to TheoryDataFun interface;
wenzelm
parents: 5001
diff changeset
   276
end;
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   277
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   278
structure GlobalSimpset = TheoryDataFun(GlobalSimpsetArgs);
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   279
val print_simpset = GlobalSimpset.print;
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   280
val simpset_ref_of_sg = GlobalSimpset.get_sg;
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   281
val simpset_ref_of = GlobalSimpset.get;
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   282
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   283
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   284
(* access global simpset *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   286
val simpset_of_sg = ! o simpset_ref_of_sg;
6391
0da748358eff Theory.sign_of;
wenzelm
parents: 6096
diff changeset
   287
val simpset_of = simpset_of_sg o Theory.sign_of;
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   288
6391
0da748358eff Theory.sign_of;
wenzelm
parents: 6096
diff changeset
   289
fun SIMPSET tacf state = tacf (simpset_of_sg (Thm.sign_of_thm state)) state;
0da748358eff Theory.sign_of;
wenzelm
parents: 6096
diff changeset
   290
fun SIMPSET' tacf i state = tacf (simpset_of_sg (Thm.sign_of_thm state)) i state;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   291
5028
61c10aad3d71 Context.the_context;
wenzelm
parents: 5024
diff changeset
   292
val simpset = simpset_of o Context.the_context;
6391
0da748358eff Theory.sign_of;
wenzelm
parents: 6096
diff changeset
   293
val simpset_ref = simpset_ref_of_sg o Theory.sign_of o Context.the_context;
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   294
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   295
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   296
(* change global simpset *)
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   297
7177
6bb7ad30f3da change_simpset_of;
wenzelm
parents: 7132
diff changeset
   298
fun change_simpset_of f x thy =
6bb7ad30f3da change_simpset_of;
wenzelm
parents: 7132
diff changeset
   299
  let val r = simpset_ref_of thy
6bb7ad30f3da change_simpset_of;
wenzelm
parents: 7132
diff changeset
   300
  in r := f (! r, x); thy end;
6bb7ad30f3da change_simpset_of;
wenzelm
parents: 7132
diff changeset
   301
6bb7ad30f3da change_simpset_of;
wenzelm
parents: 7132
diff changeset
   302
fun change_simpset f x = (change_simpset_of f x (Context.the_context ()); ());
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   303
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   304
val Addsimps = change_simpset (op addsimps);
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   305
val Delsimps = change_simpset (op delsimps);
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   306
val Addsimprocs = change_simpset (op addsimprocs);
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   307
val Delsimprocs = change_simpset (op delsimprocs);
2509
0a7169d89b7a added termless parameter;
wenzelm
parents: 2503
diff changeset
   308
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   310
(* proof data kind 'Provers/simpset' *)
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   311
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   312
structure LocalSimpsetArgs =
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   313
struct
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   314
  val name = "Provers/simpset";
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   315
  type T = simpset;
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   316
  val init = simpset_of;
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   317
  fun print _ ss = print_ss ss;
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   318
end;
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   319
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   320
structure LocalSimpset = ProofDataFun(LocalSimpsetArgs);
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   321
val print_local_simpset = LocalSimpset.print;
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   322
val get_local_simpset = LocalSimpset.get;
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   323
val put_local_simpset = LocalSimpset.put;
7273
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   324
fun map_local_simpset f ctxt = put_local_simpset (f (get_local_simpset ctxt)) ctxt;
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   325
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   326
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   327
(* attributes *)
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   328
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   329
fun change_global_ss f (thy, th) =
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   330
  let val r = simpset_ref_of thy
6096
3451f9e88528 eliminated tthm type and Attribute structure;
wenzelm
parents: 5928
diff changeset
   331
  in r := f (! r, [th]); (thy, th) end;
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   332
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   333
fun change_local_ss f (ctxt, th) =
6096
3451f9e88528 eliminated tthm type and Attribute structure;
wenzelm
parents: 5928
diff changeset
   334
  let val ss = f (get_local_simpset ctxt, [th])
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   335
  in (put_local_simpset ss ctxt, th) end;
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   336
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   337
val simp_add_global = change_global_ss (op addsimps);
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   338
val simp_del_global = change_global_ss (op delsimps);
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   339
val simp_add_local = change_local_ss (op addsimps);
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   340
val simp_del_local = change_local_ss (op delsimps);
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   341
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   342
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   343
2503
7590fd5ce3c7 cleaned up (real changes next time);
wenzelm
parents: 1770
diff changeset
   344
(** simplification tactics **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   345
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   346
fun solve_all_tac (subgoal_tac, loop_tacs, unsafe_finish_tac) mss =
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   347
  let
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   348
    val ss =
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   349
      make_ss (mss, subgoal_tac, loop_tacs,unsafe_finish_tac,unsafe_finish_tac);
4612
26764de50c74 Used THEN_ALL_NEW.
nipkow
parents: 4366
diff changeset
   350
    val solve1_tac = (subgoal_tac ss THEN_ALL_NEW (K no_tac)) 1
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   351
  in DEPTH_SOLVE solve1_tac end;
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   352
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   353
fun loop_tac loop_tacs = FIRST'(map snd loop_tacs);
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   354
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   355
(*unsafe: may instantiate unknowns that appear also in other subgoals*)
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   356
fun basic_gen_simp_tac mode finish_tac =
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   357
  fn (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, ...}) =>
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   358
    let
4612
26764de50c74 Used THEN_ALL_NEW.
nipkow
parents: 4366
diff changeset
   359
      fun simp_loop_tac i =
26764de50c74 Used THEN_ALL_NEW.
nipkow
parents: 4366
diff changeset
   360
        asm_rewrite_goal_tac mode
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   361
          (solve_all_tac (subgoal_tac,loop_tacs,unsafe_finish_tac))
4612
26764de50c74 Used THEN_ALL_NEW.
nipkow
parents: 4366
diff changeset
   362
          mss i
26764de50c74 Used THEN_ALL_NEW.
nipkow
parents: 4366
diff changeset
   363
        THEN (finish_tac (prems_of_mss mss) i ORELSE
4668
131989b78417 Little reorganization. Loop tactics have names now.
nipkow
parents: 4612
diff changeset
   364
              TRY ((loop_tac loop_tacs THEN_ALL_NEW simp_loop_tac) i))
4612
26764de50c74 Used THEN_ALL_NEW.
nipkow
parents: 4366
diff changeset
   365
    in simp_loop_tac end;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   366
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   367
fun gen_simp_tac mode 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   368
      (ss as Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, ...}) =
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   369
  basic_gen_simp_tac mode unsafe_finish_tac ss;
3551
7c013a617813 added print_ss;
wenzelm
parents: 3537
diff changeset
   370
4713
bea2ab2e360b New simplifier flag for mutual simplification.
nipkow
parents: 4682
diff changeset
   371
val          simp_tac = gen_simp_tac (false, false, false);
bea2ab2e360b New simplifier flag for mutual simplification.
nipkow
parents: 4682
diff changeset
   372
val      asm_simp_tac = gen_simp_tac (false, true, false);
bea2ab2e360b New simplifier flag for mutual simplification.
nipkow
parents: 4682
diff changeset
   373
val     full_simp_tac = gen_simp_tac (true,  false, false);
4722
d2e44673c21e The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents: 4713
diff changeset
   374
val   asm_lr_simp_tac = gen_simp_tac (true,  true, false);
d2e44673c21e The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents: 4713
diff changeset
   375
val asm_full_simp_tac = gen_simp_tac (true,  true, true);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   376
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
   377
(*not totally safe: may instantiate unknowns that appear also in other subgoals*)
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   378
fun safe_asm_full_simp_tac (ss as Simpset {mss, subgoal_tac, loop_tacs, 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   379
	unsafe_finish_tac, finish_tac}) = 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   380
  basic_gen_simp_tac (true, true, true) finish_tac ss;
2629
b442786d4469 added deleqcongs, richer rep_ss
oheimb
parents: 2567
diff changeset
   381
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   382
(*the abstraction over the proof state delays the dereferencing*)
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   383
fun          Simp_tac i st =          simp_tac (simpset ()) i st;
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   384
fun      Asm_simp_tac i st =      asm_simp_tac (simpset ()) i st;
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   385
fun     Full_simp_tac i st =     full_simp_tac (simpset ()) i st;
4722
d2e44673c21e The new asm_lr_simp_tac is the old asm_full_simp_tac.
nipkow
parents: 4713
diff changeset
   386
fun   Asm_lr_simp_tac i st =   asm_lr_simp_tac (simpset ()) i st;
4080
7dce11095b0a new implicit simpset mechanism based on Sign.sg anytype data;
wenzelm
parents: 3728
diff changeset
   387
fun Asm_full_simp_tac i st = asm_full_simp_tac (simpset ()) i st;
406
4d4e0442b106 simpset is hidden in a functor now.
nipkow
parents: 217
diff changeset
   388
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   389
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   390
5082
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   391
(** simplification rules and conversions **)
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   392
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   393
fun simp rew mode 
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   394
     (Simpset {mss, subgoal_tac, loop_tacs, unsafe_finish_tac, ...}) thm =
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   395
  let
7214
381d6987f68d exchanged finish_tac and unsafe_finish_tac (the more important one) in simpset
oheimb
parents: 7177
diff changeset
   396
    val tacf = solve_all_tac (subgoal_tac, loop_tacs, unsafe_finish_tac);
4271
3a82492e70c5 changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents: 4259
diff changeset
   397
    fun prover m th = apsome fst (Seq.pull (tacf m th));
5082
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   398
  in rew mode prover mss thm end;
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   399
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   400
val simp_thm = simp Drule.rewrite_thm;
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   401
val simp_cterm = simp Drule.rewrite_cterm;
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   402
5082
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   403
val          simplify = simp_thm (false, false, false);
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   404
val      asm_simplify = simp_thm (false, true, false);
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   405
val     full_simplify = simp_thm (true, false, false);
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   406
val asm_full_simplify = simp_thm (true, true, false);
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   407
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   408
val          rewrite = simp_cterm (false, false, false);
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   409
val      asm_rewrite = simp_cterm (false, true, false);
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   410
val     full_rewrite = simp_cterm (true, false, false);
e03460289797 added XX_YY_rewrite: simpset -> cterm -> thm;
wenzelm
parents: 5028
diff changeset
   411
val asm_full_rewrite = simp_cterm (true, true, false);
3557
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   412
9546f8185c43 added simplification meta rules;
wenzelm
parents: 3551
diff changeset
   413
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   414
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   415
(** concrete syntax of attributes **)
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   416
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   417
(* add / del *)
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   418
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   419
val simpN = "simp";
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   420
val asm_simpN = "asm_simp";
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   421
val addN = "add";
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   422
val delN = "del";
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   423
val onlyN = "only";
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   424
val otherN = "other";
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   425
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   426
fun simp_att change =
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   427
  (Args.$$$ addN >> K (op addsimps) ||
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   428
    Args.$$$ delN >> K (op delsimps) ||
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   429
    Scan.succeed (op addsimps))
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   430
  >> change
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   431
  |> Scan.lift
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   432
  |> Attrib.syntax;
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   433
5886
0373323180f5 all modifiers turned into attributes;
wenzelm
parents: 5842
diff changeset
   434
val simp_attr = (simp_att change_global_ss, simp_att change_local_ss);
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   435
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   436
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   437
(* conversions *)
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   438
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   439
fun conv_attr f =
6096
3451f9e88528 eliminated tthm type and Attribute structure;
wenzelm
parents: 5928
diff changeset
   440
  (Attrib.no_args (Drule.rule_attribute (f o simpset_of)),
3451f9e88528 eliminated tthm type and Attribute structure;
wenzelm
parents: 5928
diff changeset
   441
    Attrib.no_args (Drule.rule_attribute (f o get_local_simpset)));
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   442
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   443
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   444
(* setup_attrs *)
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   445
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   446
val setup_attrs = Attrib.add_attributes
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   447
 [(simpN,               simp_attr, "simplification rule"),
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   448
  ("simplify",          conv_attr simplify, "simplify rule"),
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   449
  ("asm_simplify",      conv_attr asm_simplify, "simplify rule, using assumptions"),
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   450
  ("full_simplify",     conv_attr full_simplify, "fully simplify rule"),
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   451
  ("asm_full_simplify", conv_attr asm_full_simplify, "fully simplify rule, using assumptions")];
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   452
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   453
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   454
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   455
(** proof methods **)
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   456
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   457
(* simplification *)
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   458
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   459
val simp_modifiers =
7273
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   460
 [Args.$$$ simpN -- Args.$$$ addN >> K (I, simp_add_local),
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   461
  Args.$$$ simpN -- Args.$$$ delN >> K (I, simp_del_local),
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   462
  Args.$$$ simpN -- Args.$$$ onlyN >> K (map_local_simpset clear_ss, simp_add_local),
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   463
  Args.$$$ otherN >> K (I, I)];
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   464
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   465
val simp_modifiers' =
7273
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   466
 [Args.$$$ addN >> K (I, simp_add_local),
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   467
  Args.$$$ delN >> K (I, simp_del_local),
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   468
  Args.$$$ onlyN >> K (map_local_simpset clear_ss, simp_add_local),
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   469
  Args.$$$ otherN >> K (I, I)];
5928
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   470
6e00a206a948 export simp_modifiers;
wenzelm
parents: 5886
diff changeset
   471
val simp_args = Method.only_sectioned_args simp_modifiers';
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   472
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   473
fun simp_meth thin cut tac ctxt = Method.METHOD (fn facts =>
6534
5a838c1d9d2f improper simp methods;
wenzelm
parents: 6391
diff changeset
   474
  FIRSTGOAL
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   475
    ((if thin then REPEAT_DETERM o etac Drule.thin_rl else K all_tac)
7423
28b48fcb0d55 Method.insert_tac;
wenzelm
parents: 7273
diff changeset
   476
      THEN' (if cut then Method.insert_tac facts else K all_tac)
6534
5a838c1d9d2f improper simp methods;
wenzelm
parents: 6391
diff changeset
   477
      THEN' tac (get_local_simpset ctxt)));
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   478
6911
ef0f25d0bc2d simp only: attribute, method arg;
wenzelm
parents: 6556
diff changeset
   479
val simp_method = simp_args ooo simp_meth;
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   480
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   481
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   482
(* setup_methods *)
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   483
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   484
val setup_methods = Method.add_methods
7423
28b48fcb0d55 Method.insert_tac;
wenzelm
parents: 7273
diff changeset
   485
 [(simpN,               simp_method true true (CHANGED oo asm_full_simp_tac),
28b48fcb0d55 Method.insert_tac;
wenzelm
parents: 7273
diff changeset
   486
    "full simplification (including facts, excluding assumptions)"),
7273
d80b9be87535 Method.modifier;
wenzelm
parents: 7217
diff changeset
   487
  (asm_simpN,           simp_method false true (CHANGED oo asm_full_simp_tac),
7423
28b48fcb0d55 Method.insert_tac;
wenzelm
parents: 7273
diff changeset
   488
    "full simplification (including facts and assumptions)"),
7132
5c31d06ead04 eliminated METHOD0 in favour of same_tac;
wenzelm
parents: 6994
diff changeset
   489
  ("simp_tac",          simp_method false false simp_tac, "simp_tac (improper!)"),
5c31d06ead04 eliminated METHOD0 in favour of same_tac;
wenzelm
parents: 6994
diff changeset
   490
  ("asm_simp_tac",      simp_method false false asm_simp_tac, "asm_simp_tac (improper!)"),
5c31d06ead04 eliminated METHOD0 in favour of same_tac;
wenzelm
parents: 6994
diff changeset
   491
  ("full_simp_tac",     simp_method false false full_simp_tac, "full_simp_tac (improper!)"),
5c31d06ead04 eliminated METHOD0 in favour of same_tac;
wenzelm
parents: 6994
diff changeset
   492
  ("asm_full_simp_tac", simp_method false false asm_full_simp_tac, "asm_full_simp_tac (improper!)"),
5c31d06ead04 eliminated METHOD0 in favour of same_tac;
wenzelm
parents: 6994
diff changeset
   493
  ("asm_lr_simp_tac",   simp_method false false asm_lr_simp_tac, "asm_lr_simp_tac (improper!)")];
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   494
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   495
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   496
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   497
(** theory setup **)
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   498
5842
1a708aa63ff0 local simpset theory data;
wenzelm
parents: 5082
diff changeset
   499
val setup = [GlobalSimpset.init, LocalSimpset.init, setup_attrs, setup_methods];
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   500
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   501
1243
fa09705a5890 added global simpset
clasohm
parents: 983
diff changeset
   502
end;
4795
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   503
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   504
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   505
structure BasicSimplifier: BASIC_SIMPLIFIER = Simplifier;
721b532ada7a tuned comments;
wenzelm
parents: 4784
diff changeset
   506
open BasicSimplifier;