src/Provers/clasimp.ML
author wenzelm
Sun, 15 Mar 2009 15:59:44 +0100
changeset 30528 7173bf123335
parent 30513 1796b8ea88aa
child 30541 9f168bdc468a
permissions -rw-r--r--
simplified attribute setup;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
     1
(*  Title:      Provers/clasimp.ML
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
     2
    Author:     David von Oheimb, TU Muenchen
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
     3
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
     4
Combination of classical reasoner and simplifier (depends on
16019
0e1405402d53 Simplifier already setup in Pure;
wenzelm
parents: 15570
diff changeset
     5
splitter.ML, classical.ML, blast.ML).
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
     6
*)
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
     7
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
     8
infix 4 addSIs2 addSEs2 addSDs2 addIs2 addEs2 addDs2 addsimps2 delsimps2
13603
57f364d1d3b2 Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
berghofe
parents: 13026
diff changeset
     9
  addSss addss addss' addIffs delIffs;
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    10
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    11
signature CLASIMP_DATA =
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    12
sig
8469
482c301041b4 include Splitter.split_modifiers;
wenzelm
parents: 8168
diff changeset
    13
  structure Splitter: SPLITTER
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    14
  structure Classical: CLASSICAL
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    15
  structure Blast: BLAST
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    16
  sharing type Classical.claset = Blast.claset
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    17
  val notE: thm
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    18
  val iffD1: thm
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    19
  val iffD2: thm
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    20
end
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    21
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    22
signature CLASIMP =
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    23
sig
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    24
  include CLASIMP_DATA
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    25
  type claset
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    26
  type clasimpset
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    27
  val get_css: Context.generic -> clasimpset
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    28
  val map_css: (clasimpset -> clasimpset) -> Context.generic -> Context.generic
17879
88844eea4ce2 functor: no Simplifier argument;
wenzelm
parents: 17084
diff changeset
    29
  val change_clasimpset: (clasimpset -> clasimpset) -> unit
10317
3205fe2f4ef5 added clasimpset: unit -> clasimpset;
wenzelm
parents: 10036
diff changeset
    30
  val clasimpset: unit -> clasimpset
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    31
  val local_clasimpset_of: Proof.context -> clasimpset
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    32
  val addSIs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    33
  val addSEs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    34
  val addSDs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    35
  val addIs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    36
  val addEs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    37
  val addDs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    38
  val addsimps2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    39
  val delsimps2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    40
  val addSss: claset * simpset -> claset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    41
  val addss: claset * simpset -> claset
13603
57f364d1d3b2 Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
berghofe
parents: 13026
diff changeset
    42
  val addss': claset * simpset -> claset
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    43
  val addIffs: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    44
  val delIffs: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    45
  val AddIffs: thm list -> unit
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    46
  val DelIffs: thm list -> unit
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    47
  val CLASIMPSET: (clasimpset -> tactic) -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    48
  val CLASIMPSET': (clasimpset -> 'a -> tactic) -> 'a -> tactic
5483
2fc3f4450fe8 added clarsimp_tac and Clarsimp_tac
oheimb
parents: 5219
diff changeset
    49
  val clarsimp_tac: clasimpset -> int -> tactic
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    50
  val Clarsimp_tac: int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    51
  val mk_auto_tac: clasimpset -> int -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    52
  val auto_tac: clasimpset -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    53
  val Auto_tac: tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    54
  val auto: unit -> unit
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    55
  val force_tac: clasimpset  -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    56
  val Force_tac: int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    57
  val force: int -> unit
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    58
  val fast_simp_tac: clasimpset -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    59
  val slow_simp_tac: clasimpset -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    60
  val best_simp_tac: clasimpset -> int -> tactic
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    61
  val attrib: (clasimpset * thm list -> clasimpset) -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    62
  val iff_add: attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    63
  val iff_add': attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    64
  val iff_del: attribute
30513
1796b8ea88aa eliminated type Args.T;
wenzelm
parents: 30510
diff changeset
    65
  val iff_modifiers: Method.modifier parser list
1796b8ea88aa eliminated type Args.T;
wenzelm
parents: 30510
diff changeset
    66
  val clasimp_modifiers: Method.modifier parser list
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    67
  val clasimp_setup: theory -> theory
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    68
end;
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    69
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    70
functor ClasimpFun(Data: CLASIMP_DATA): CLASIMP =
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    71
struct
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    72
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    73
open Data;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    74
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    75
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    76
(* type clasimpset *)
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    77
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    78
type claset = Classical.claset;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    79
type clasimpset = claset * simpset;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    80
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    81
fun get_css context = (Classical.get_cs context, Simplifier.get_ss context);
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    82
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    83
fun map_css f context =
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    84
  let val (cs', ss') = f (get_css context)
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    85
  in context |> Classical.map_cs (K cs') |> Simplifier.map_ss (K ss') end;
17879
88844eea4ce2 functor: no Simplifier argument;
wenzelm
parents: 17084
diff changeset
    86
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    87
fun change_clasimpset f = Context.>> (fn context => (Context.the_theory context; map_css f context));
10317
3205fe2f4ef5 added clasimpset: unit -> clasimpset;
wenzelm
parents: 10036
diff changeset
    88
fun clasimpset () = (Classical.claset (), Simplifier.simpset ());
3205fe2f4ef5 added clasimpset: unit -> clasimpset;
wenzelm
parents: 10036
diff changeset
    89
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    90
fun local_clasimpset_of ctxt =
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    91
  (Classical.local_claset_of ctxt, Simplifier.local_simpset_of ctxt);
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    92
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    93
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    94
(* clasimpset operations *)
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    95
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    96
(*this interface for updating a clasimpset is rudimentary and just for
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    97
  convenience for the most common cases*)
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    98
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    99
fun pair_upd1 f ((a, b), x) = (f (a, x), b);
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   100
fun pair_upd2 f ((a, b), x) = (a, f (b, x));
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   101
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   102
fun op addSIs2   arg = pair_upd1 Classical.addSIs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   103
fun op addSEs2   arg = pair_upd1 Classical.addSEs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   104
fun op addSDs2   arg = pair_upd1 Classical.addSDs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   105
fun op addIs2    arg = pair_upd1 Classical.addIs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   106
fun op addEs2    arg = pair_upd1 Classical.addEs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   107
fun op addDs2    arg = pair_upd1 Classical.addDs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   108
fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   109
fun op delsimps2 arg = pair_upd2 Simplifier.delsimps arg;
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   110
9402
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 8639
diff changeset
   111
(*not totally safe: may instantiate unknowns that appear also in other subgoals*)
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 8639
diff changeset
   112
val safe_asm_full_simp_tac = Simplifier.generic_simp_tac true (true,true,true);
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 8639
diff changeset
   113
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   114
(*Add a simpset to a classical set!*)
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   115
(*Caution: only one simpset added can be added by each of addSss and addss*)
11181
d04f57b91166 renamed addaltern to addafter, addSaltern to addSafter
oheimb
parents: 10821
diff changeset
   116
fun cs addSss ss = Classical.addSafter (cs, ("safe_asm_full_simp_tac",
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   117
                            CHANGED o safe_asm_full_simp_tac ss));
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   118
fun cs addss  ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   119
                            CHANGED o Simplifier.asm_full_simp_tac ss));
13603
57f364d1d3b2 Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
berghofe
parents: 13026
diff changeset
   120
fun cs addss' ss = Classical.addbefore  (cs, ("asm_full_simp_tac",
57f364d1d3b2 Introduced addss', which adds asm_lr_simp_tac as a wrapper to the claset.
berghofe
parents: 13026
diff changeset
   121
                            CHANGED o Simplifier.asm_lr_simp_tac ss));
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   122
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   123
(* iffs: addition of rules to simpsets and clasets simultaneously *)
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   124
11344
57b7ad51971c streamlined addIffs/delIffs, added warnings
oheimb
parents: 11181
diff changeset
   125
(*Takes (possibly conditional) theorems of the form A<->B to
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   126
        the Safe Intr     rule B==>A and
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   127
        the Safe Destruct rule A==>B.
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   128
  Also ~A goes to the Safe Elim rule A ==> ?R
11462
cf3e7f5ad0e1 removed the warning from [iff]
paulson
parents: 11344
diff changeset
   129
  Failing other cases, A is added as a Safe Intr rule*)
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   130
local
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   131
13026
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   132
fun addIff decls1 decls2 simp ((cs, ss), th) =
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   133
  let
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   134
    val n = nprems_of th;
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 16019
diff changeset
   135
    val (elim, intro) = if n = 0 then decls1 else decls2;
13026
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   136
    val zero_rotate = zero_var_indexes o rotate_prems n;
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   137
  in
21709
9cfd9eb9faec removed use of put_name_hint, as the ATP linkup no longer needs this
paulson
parents: 21646
diff changeset
   138
    (elim (intro (cs, [zero_rotate (th RS iffD2)]),
9cfd9eb9faec removed use of put_name_hint, as the ATP linkup no longer needs this
paulson
parents: 21646
diff changeset
   139
           [Tactic.make_elim (zero_rotate (th RS iffD1))])
9cfd9eb9faec removed use of put_name_hint, as the ATP linkup no longer needs this
paulson
parents: 21646
diff changeset
   140
      handle THM _ => (elim (cs, [zero_rotate (th RS notE)])
11902
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   141
        handle THM _ => intro (cs, [th])),
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   142
     simp (ss, [th]))
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   143
  end;
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   144
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   145
fun delIff delcs delss ((cs, ss), th) =
11902
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   146
  let val zero_rotate = zero_var_indexes o rotate_prems (nprems_of th) in
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   147
    (delcs (cs, [zero_rotate (th RS iffD2),
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   148
        Tactic.make_elim (zero_rotate (th RS iffD1))])
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   149
      handle THM _ => (delcs (cs, [zero_rotate (th RS notE)])
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   150
        handle THM _ => delcs (cs, [th])),
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   151
     delss (ss, [th]))
11902
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   152
  end;
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   153
18630
69fe387b3b6e generic attributes;
wenzelm
parents: 18529
diff changeset
   154
fun modifier att (x, ths) =
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 27809
diff changeset
   155
  fst (Library.foldl_map (Library.apply [att]) (x, rev ths));
18630
69fe387b3b6e generic attributes;
wenzelm
parents: 18529
diff changeset
   156
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   157
val addXIs = modifier (ContextRules.intro_query NONE);
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   158
val addXEs = modifier (ContextRules.elim_query NONE);
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   159
val addXDs = modifier (ContextRules.dest_query NONE);
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   160
val delXs = modifier ContextRules.rule_del;
18630
69fe387b3b6e generic attributes;
wenzelm
parents: 18529
diff changeset
   161
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   162
in
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   163
10033
fc4e7432b2b1 added iff_add_global', iff_add_local' (syntax "iff?");
wenzelm
parents: 9952
diff changeset
   164
val op addIffs =
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   165
  Library.foldl
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 16019
diff changeset
   166
      (addIff (Classical.addSEs, Classical.addSIs)
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   167
              (Classical.addEs, Classical.addIs)
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 16019
diff changeset
   168
              Simplifier.addsimps);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   169
val op delIffs = Library.foldl (delIff Classical.delrules Simplifier.delsimps);
10033
fc4e7432b2b1 added iff_add_global', iff_add_local' (syntax "iff?");
wenzelm
parents: 9952
diff changeset
   170
17879
88844eea4ce2 functor: no Simplifier argument;
wenzelm
parents: 17084
diff changeset
   171
fun AddIffs thms = change_clasimpset (fn css => css addIffs thms);
88844eea4ce2 functor: no Simplifier argument;
wenzelm
parents: 17084
diff changeset
   172
fun DelIffs thms = change_clasimpset (fn css => css delIffs thms);
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   173
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   174
fun addIffs_generic (x, ths) =
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   175
  Library.foldl (addIff (addXEs, addXIs) (addXEs, addXIs) #1) ((x, ()), ths) |> #1;
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   176
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   177
fun delIffs_generic (x, ths) =
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   178
  Library.foldl (delIff delXs #1) ((x, ()), ths) |> #1;
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   179
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   180
end;
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   181
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   182
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   183
(* tacticals *)
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   184
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   185
fun CLASIMPSET tacf state =
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   186
  Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss))) state;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   187
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   188
fun CLASIMPSET' tacf i state =
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   189
  Classical.CLASET (fn cs => Simplifier.SIMPSET (fn ss => tacf (cs, ss) i)) state;
4888
7301ff9f412b added CLASIMPSET(') tacticals;
wenzelm
parents: 4884
diff changeset
   190
7301ff9f412b added CLASIMPSET(') tacticals;
wenzelm
parents: 4884
diff changeset
   191
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   192
fun clarsimp_tac (cs, ss) =
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   193
  safe_asm_full_simp_tac ss THEN_ALL_NEW
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   194
  Classical.clarify_tac (cs addSss ss);
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   195
10317
3205fe2f4ef5 added clasimpset: unit -> clasimpset;
wenzelm
parents: 10036
diff changeset
   196
fun Clarsimp_tac i = clarsimp_tac (clasimpset ()) i;
5483
2fc3f4450fe8 added clarsimp_tac and Clarsimp_tac
oheimb
parents: 5219
diff changeset
   197
2fc3f4450fe8 added clarsimp_tac and Clarsimp_tac
oheimb
parents: 5219
diff changeset
   198
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   199
(* auto_tac *)
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   200
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   201
fun blast_depth_tac cs m i thm =
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   202
    Blast.depth_tac cs m i thm
5554
3cae5d6510c2 removed addcongs2 and delcongs2
oheimb
parents: 5525
diff changeset
   203
      handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty);
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   204
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   205
(* a variant of depth_tac that avoids interference of the simplifier
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   206
   with dup_step_tac when they are combined by auto_tac *)
5756
8ef5288c24b0 corrected auto_tac (applications of unsafe wrappers)
oheimb
parents: 5567
diff changeset
   207
local
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   208
fun slow_step_tac' cs = Classical.appWrappers cs
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   209
        (Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs);
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   210
in fun nodup_depth_tac cs m i state = SELECT_GOAL
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   211
   (Classical.safe_steps_tac cs 1 THEN_ELSE
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   212
        (DEPTH_SOLVE (nodup_depth_tac cs m 1),
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   213
         Classical.inst0_step_tac cs 1 APPEND COND (K (m=0)) no_tac
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   214
             (slow_step_tac' cs 1 THEN DEPTH_SOLVE (nodup_depth_tac cs (m-1) 1))
5756
8ef5288c24b0 corrected auto_tac (applications of unsafe wrappers)
oheimb
parents: 5567
diff changeset
   215
        )) i state;
8ef5288c24b0 corrected auto_tac (applications of unsafe wrappers)
oheimb
parents: 5567
diff changeset
   216
end;
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   217
9402
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 8639
diff changeset
   218
(*Designed to be idempotent, except if blast_depth_tac instantiates variables
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   219
  in some of the subgoals*)
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   220
fun mk_auto_tac (cs, ss) m n =
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   221
    let val cs' = cs addss ss
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   222
        val maintac =
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   223
          blast_depth_tac cs m               (* fast but can't use wrappers *)
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   224
          ORELSE'
5756
8ef5288c24b0 corrected auto_tac (applications of unsafe wrappers)
oheimb
parents: 5567
diff changeset
   225
          (CHANGED o nodup_depth_tac cs' n); (* slower but more general *)
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   226
    in  EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss),
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   227
               TRY (Classical.safe_tac cs),
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   228
               REPEAT (FIRSTGOAL maintac),
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   229
               TRY (Classical.safe_tac (cs addSss ss)),
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   230
               prune_params_tac]
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   231
    end;
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   232
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   233
fun auto_tac css = mk_auto_tac css 4 2;
10317
3205fe2f4ef5 added clasimpset: unit -> clasimpset;
wenzelm
parents: 10036
diff changeset
   234
fun Auto_tac st = auto_tac (clasimpset ()) st;
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   235
fun auto () = by Auto_tac;
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   236
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   237
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   238
(* force_tac *)
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   239
4659
a78ecc7341e3 added smart_tac
oheimb
parents: 4652
diff changeset
   240
(* aimed to solve the given subgoal totally, using whatever tools possible *)
4717
1370ad043564 renamed smart_tac to force_tac, slight improvement of force_tac
oheimb
parents: 4659
diff changeset
   241
fun force_tac (cs,ss) = let val cs' = cs addss ss in SELECT_GOAL (EVERY [
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   242
        Classical.clarify_tac cs' 1,
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   243
        IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1),
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   244
        ALLGOALS (Classical.first_best_tac cs')]) end;
10317
3205fe2f4ef5 added clasimpset: unit -> clasimpset;
wenzelm
parents: 10036
diff changeset
   245
fun Force_tac i = force_tac (clasimpset ()) i;
4717
1370ad043564 renamed smart_tac to force_tac, slight improvement of force_tac
oheimb
parents: 4659
diff changeset
   246
fun force i = by (Force_tac i);
4659
a78ecc7341e3 added smart_tac
oheimb
parents: 4652
diff changeset
   247
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   248
9805
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   249
(* basic combinations *)
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   250
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   251
fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end;
9591
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   252
9805
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   253
val fast_simp_tac = ADDSS Classical.fast_tac;
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   254
val slow_simp_tac = ADDSS Classical.slow_tac;
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   255
val best_simp_tac = ADDSS Classical.best_tac;
9591
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   256
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   257
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   258
(* access clasimpset *)
8639
31bcb6b64d60 added change_global/local_css;
wenzelm
parents: 8469
diff changeset
   259
15032
02aed07e01bf local_cla/simpset_of;
wenzelm
parents: 14981
diff changeset
   260
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   261
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   262
(* attributes *)
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   263
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   264
fun attrib f = Thm.declaration_attribute (fn th => map_css (fn css => f (css, [th])));
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   265
fun attrib' f (x, th) = (f (x, [th]), th);
10033
fc4e7432b2b1 added iff_add_global', iff_add_local' (syntax "iff?");
wenzelm
parents: 9952
diff changeset
   266
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   267
val iff_add = attrib (op addIffs);
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   268
val iff_add' = attrib' addIffs_generic;
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   269
val iff_del = attrib (op delIffs) o attrib' delIffs_generic;
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   270
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30513
diff changeset
   271
fun iff_att x = (Scan.lift
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   272
 (Args.del >> K iff_del ||
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   273
  Scan.option Args.add -- Args.query >> K iff_add' ||
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30513
diff changeset
   274
  Scan.option Args.add >> K iff_add)) x;
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   275
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   276
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   277
(* method modifiers *)
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   278
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   279
val iffN = "iff";
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   280
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   281
val iff_modifiers =
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   282
 [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add): Method.modifier),
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   283
  Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add'),
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   284
  Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del)];
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   285
8469
482c301041b4 include Splitter.split_modifiers;
wenzelm
parents: 8168
diff changeset
   286
val clasimp_modifiers =
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   287
  Simplifier.simp_modifiers @ Splitter.split_modifiers @
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   288
  Classical.cla_modifiers @ iff_modifiers;
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   289
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   290
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   291
(* methods *)
5926
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   292
30510
4120fc59dd85 unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents: 30190
diff changeset
   293
fun clasimp_meth tac prems ctxt = METHOD (fn facts =>
15032
02aed07e01bf local_cla/simpset_of;
wenzelm
parents: 14981
diff changeset
   294
  ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_clasimpset_of ctxt));
7132
5c31d06ead04 eliminated METHOD0 in favour of same_tac;
wenzelm
parents: 5985
diff changeset
   295
30510
4120fc59dd85 unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents: 30190
diff changeset
   296
fun clasimp_meth' tac prems ctxt = METHOD (fn facts =>
15032
02aed07e01bf local_cla/simpset_of;
wenzelm
parents: 14981
diff changeset
   297
  HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_clasimpset_of ctxt)));
5926
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   298
7559
1d2c099e98f7 setup for refined facts handling;
wenzelm
parents: 7437
diff changeset
   299
val clasimp_method = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth;
1d2c099e98f7 setup for refined facts handling;
wenzelm
parents: 7437
diff changeset
   300
val clasimp_method' = Method.bang_sectioned_args clasimp_modifiers o clasimp_meth';
5926
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   301
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   302
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   303
fun auto_args m =
27809
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 26497
diff changeset
   304
  Method.bang_sectioned_args' clasimp_modifiers
a1e409db516b unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents: 26497
diff changeset
   305
    (Scan.lift (Scan.option (OuterParse.nat -- OuterParse.nat))) m;
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   306
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15032
diff changeset
   307
fun auto_meth NONE = clasimp_meth (CHANGED_PROP o auto_tac)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15032
diff changeset
   308
  | auto_meth (SOME (m, n)) = clasimp_meth (CHANGED_PROP o (fn css => mk_auto_tac css m n));
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   309
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   310
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   311
(* theory setup *)
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   312
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   313
val clasimp_setup =
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30513
diff changeset
   314
 (Attrib.setup @{binding iff} iff_att "declaration of Simplifier / Classical rules" #>
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   315
  Method.add_methods
9591
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   316
   [("fastsimp", clasimp_method' fast_simp_tac, "combined fast and simp"),
9805
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   317
    ("slowsimp", clasimp_method' slow_simp_tac, "combined slow and simp"),
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   318
    ("bestsimp", clasimp_method' best_simp_tac, "combined best and simp"),
9591
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   319
    ("force", clasimp_method' force_tac, "force"),
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   320
    ("auto", auto_args auto_meth, "auto"),
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18688
diff changeset
   321
    ("clarsimp", clasimp_method' (CHANGED_PROP oo clarsimp_tac), "clarify simplified goal")]);
5926
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   322
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   323
end;