src/Provers/clasimp.ML
author wenzelm
Tue, 26 Apr 2011 21:49:39 +0200
changeset 42478 8a526c010c3b
parent 42373 fb539d6d1ac2
child 42479 b7c9f09d4d88
permissions -rw-r--r--
modernized Clasimp 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
  type claset
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    25
  type clasimpset
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    26
  val get_css: Context.generic -> clasimpset
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    27
  val map_css: (clasimpset -> clasimpset) -> Context.generic -> Context.generic
32148
253f6808dabe renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents: 30609
diff changeset
    28
  val clasimpset_of: Proof.context -> clasimpset
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    29
  val addSIs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    30
  val addSEs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    31
  val addSDs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    32
  val addIs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    33
  val addEs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    34
  val addDs2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    35
  val addsimps2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    36
  val delsimps2: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    37
  val addSss: claset * simpset -> claset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    38
  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
    39
  val addss': claset * simpset -> claset
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    40
  val addIffs: clasimpset * thm list -> clasimpset
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    41
  val delIffs: clasimpset * thm list -> clasimpset
5483
2fc3f4450fe8 added clarsimp_tac and Clarsimp_tac
oheimb
parents: 5219
diff changeset
    42
  val clarsimp_tac: clasimpset -> int -> tactic
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    43
  val mk_auto_tac: clasimpset -> int -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    44
  val auto_tac: clasimpset -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    45
  val force_tac: clasimpset  -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    46
  val fast_simp_tac: clasimpset -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    47
  val slow_simp_tac: clasimpset -> int -> tactic
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
    48
  val best_simp_tac: clasimpset -> int -> tactic
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    49
  val attrib: (clasimpset * thm list -> clasimpset) -> attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    50
  val iff_add: attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    51
  val iff_add': attribute
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    52
  val iff_del: attribute
30513
1796b8ea88aa eliminated type Args.T;
wenzelm
parents: 30510
diff changeset
    53
  val iff_modifiers: Method.modifier parser list
1796b8ea88aa eliminated type Args.T;
wenzelm
parents: 30510
diff changeset
    54
  val clasimp_modifiers: Method.modifier parser list
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    55
  val clasimp_setup: theory -> theory
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    56
end;
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    57
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
    58
functor Clasimp(Data: CLASIMP_DATA): CLASIMP =
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    59
struct
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
    60
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
    61
structure Splitter = Data.Splitter;
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
    62
structure Classical = Data.Classical;
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
    63
structure Blast = Data.Blast;
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    64
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    65
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    66
(* type clasimpset *)
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    67
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    68
type claset = Classical.claset;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    69
type clasimpset = claset * simpset;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    70
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    71
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
    72
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    73
fun map_css f context =
36601
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    74
  let
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    75
    val (cs, ss) = get_css context;
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    76
    val (cs', ss') = f (cs, Simplifier.context (Context.proof_of context) ss);
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    77
  in
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    78
    context
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    79
    |> Classical.map_cs (K cs')
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    80
    |> Simplifier.map_ss (K (Simplifier.inherit_context ss ss'))
8a041e2d8122 map_css: explicitly provide a temporary context, which might be required for mksimps, mkcongs etc.;
wenzelm
parents: 35613
diff changeset
    81
  end;
17879
88844eea4ce2 functor: no Simplifier argument;
wenzelm
parents: 17084
diff changeset
    82
32148
253f6808dabe renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
wenzelm
parents: 30609
diff changeset
    83
fun clasimpset_of ctxt = (Classical.claset_of ctxt, Simplifier.simpset_of ctxt);
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    84
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    85
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    86
(* clasimpset operations *)
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    87
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    88
(*this interface for updating a clasimpset is rudimentary and just for
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    89
  convenience for the most common cases*)
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    90
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    91
fun pair_upd1 f ((a, b), x) = (f (a, x), b);
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
    92
fun pair_upd2 f ((a, b), x) = (a, f (b, x));
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    93
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    94
fun op addSIs2   arg = pair_upd1 Classical.addSIs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    95
fun op addSEs2   arg = pair_upd1 Classical.addSEs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    96
fun op addSDs2   arg = pair_upd1 Classical.addSDs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    97
fun op addIs2    arg = pair_upd1 Classical.addIs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    98
fun op addEs2    arg = pair_upd1 Classical.addEs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
    99
fun op addDs2    arg = pair_upd1 Classical.addDs arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   100
fun op addsimps2 arg = pair_upd2 Simplifier.addsimps arg;
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   101
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
   102
9402
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 8639
diff changeset
   103
(*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
   104
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
   105
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   106
(*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
   107
(*Caution: only one simpset added can be added by each of addSss and addss*)
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   108
fun cs addSss ss =
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   109
  Classical.addSafter (cs, ("safe_asm_full_simp_tac", CHANGED o safe_asm_full_simp_tac ss));
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   110
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   111
fun cs addss ss =
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   112
  Classical.addbefore (cs, ("asm_full_simp_tac", CHANGED o Simplifier.asm_full_simp_tac ss));
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   113
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   114
fun cs addss' ss =
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   115
  Classical.addbefore (cs, ("asm_full_simp_tac", 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
   116
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   117
(* iffs: addition of rules to simpsets and clasets simultaneously *)
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   118
11344
57b7ad51971c streamlined addIffs/delIffs, added warnings
oheimb
parents: 11181
diff changeset
   119
(*Takes (possibly conditional) theorems of the form A<->B to
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   120
        the Safe Intr     rule B==>A and
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   121
        the Safe Destruct rule A==>B.
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   122
  Also ~A goes to the Safe Elim rule A ==> ?R
11462
cf3e7f5ad0e1 removed the warning from [iff]
paulson
parents: 11344
diff changeset
   123
  Failing other cases, A is added as a Safe Intr rule*)
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   124
local
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   125
13026
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   126
fun addIff decls1 decls2 simp ((cs, ss), th) =
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   127
  let
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   128
    val n = nprems_of th;
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 16019
diff changeset
   129
    val (elim, intro) = if n = 0 then decls1 else decls2;
13026
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   130
    val zero_rotate = zero_var_indexes o rotate_prems n;
e45ebbb2e18e iff: conditional rules declared as ``unsafe'';
wenzelm
parents: 12375
diff changeset
   131
  in
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   132
    (elim (intro (cs, [zero_rotate (th RS Data.iffD2)]),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   133
           [Tactic.make_elim (zero_rotate (th RS Data.iffD1))])
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   134
      handle THM _ => (elim (cs, [zero_rotate (th RS Data.notE)])
11902
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   135
        handle THM _ => intro (cs, [th])),
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   136
     simp (ss, [th]))
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   137
  end;
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   138
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   139
fun delIff delcs delss ((cs, ss), th) =
11902
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   140
  let val zero_rotate = zero_var_indexes o rotate_prems (nprems_of th) in
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   141
    (delcs (cs, [zero_rotate (th RS Data.iffD2),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   142
        Tactic.make_elim (zero_rotate (th RS Data.iffD1))])
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   143
      handle THM _ => (delcs (cs, [zero_rotate (th RS Data.notE)])
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   144
        handle THM _ => delcs (cs, [th])),
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   145
     delss (ss, [th]))
11902
db207d68392c iff: always rotate prems;
wenzelm
parents: 11496
diff changeset
   146
  end;
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   147
18630
69fe387b3b6e generic attributes;
wenzelm
parents: 18529
diff changeset
   148
fun modifier att (x, ths) =
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 27809
diff changeset
   149
  fst (Library.foldl_map (Library.apply [att]) (x, rev ths));
18630
69fe387b3b6e generic attributes;
wenzelm
parents: 18529
diff changeset
   150
33369
470a7b233ee5 modernized structure Context_Rules;
wenzelm
parents: 32861
diff changeset
   151
val addXIs = modifier (Context_Rules.intro_query NONE);
470a7b233ee5 modernized structure Context_Rules;
wenzelm
parents: 32861
diff changeset
   152
val addXEs = modifier (Context_Rules.elim_query NONE);
470a7b233ee5 modernized structure Context_Rules;
wenzelm
parents: 32861
diff changeset
   153
val delXs = modifier Context_Rules.rule_del;
18630
69fe387b3b6e generic attributes;
wenzelm
parents: 18529
diff changeset
   154
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   155
in
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   156
10033
fc4e7432b2b1 added iff_add_global', iff_add_local' (syntax "iff?");
wenzelm
parents: 9952
diff changeset
   157
val op addIffs =
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   158
  Library.foldl
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 16019
diff changeset
   159
      (addIff (Classical.addSEs, Classical.addSIs)
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   160
              (Classical.addEs, Classical.addIs)
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 16019
diff changeset
   161
              Simplifier.addsimps);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   162
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
   163
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   164
fun addIffs_generic (x, ths) =
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   165
  Library.foldl (addIff (addXEs, addXIs) (addXEs, addXIs) #1) ((x, ()), ths) |> #1;
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   166
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   167
fun delIffs_generic (x, ths) =
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   168
  Library.foldl (delIff delXs #1) ((x, ()), ths) |> #1;
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   169
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   170
end;
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   171
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   172
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   173
(* tactics *)
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   174
12375
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   175
fun clarsimp_tac (cs, ss) =
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   176
  safe_asm_full_simp_tac ss THEN_ALL_NEW
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   177
  Classical.clarify_tac (cs addSss ss);
539a32568db3 iff?: refer to Pure/ContextRules;
wenzelm
parents: 11902
diff changeset
   178
5483
2fc3f4450fe8 added clarsimp_tac and Clarsimp_tac
oheimb
parents: 5219
diff changeset
   179
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   180
(* auto_tac *)
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   181
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   182
fun blast_depth_tac cs m i thm =
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   183
  Blast.depth_tac cs m i thm
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   184
    handle Blast.TRANS s => (warning ("Blast_tac: " ^ s); Seq.empty);
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   185
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   186
(* a variant of depth_tac that avoids interference of the simplifier
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   187
   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
   188
local
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   189
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   190
fun slow_step_tac' cs =
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   191
  Classical.appWrappers cs
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   192
    (Classical.instp_step_tac cs APPEND' Classical.haz_step_tac cs);
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   193
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   194
in
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   195
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   196
fun nodup_depth_tac cs m i state =
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   197
  SELECT_GOAL
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   198
    (Classical.safe_steps_tac cs 1 THEN_ELSE
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   199
      (DEPTH_SOLVE (nodup_depth_tac cs m 1),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   200
        Classical.inst0_step_tac cs 1 APPEND COND (K (m = 0)) no_tac
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   201
          (slow_step_tac' cs 1 THEN DEPTH_SOLVE (nodup_depth_tac cs (m - 1) 1)))) i state;
5756
8ef5288c24b0 corrected auto_tac (applications of unsafe wrappers)
oheimb
parents: 5567
diff changeset
   202
end;
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   203
9402
480a1b40fdd6 strengthened force_tac by using new first_best_tac
oheimb
parents: 8639
diff changeset
   204
(*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
   205
  in some of the subgoals*)
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   206
fun mk_auto_tac (cs, ss) m n =
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   207
  let
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   208
    val cs' = cs addss ss;
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   209
    val main_tac =
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   210
      blast_depth_tac cs m               (* fast but can't use wrappers *)
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   211
      ORELSE'
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   212
      (CHANGED o nodup_depth_tac cs' n); (* slower but more general *)
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   213
  in
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   214
    EVERY [PARALLEL_GOALS (ALLGOALS (Simplifier.asm_full_simp_tac ss)),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   215
      TRY (Classical.safe_tac cs),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   216
      REPEAT_DETERM (FIRSTGOAL main_tac),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   217
      TRY (Classical.safe_tac (cs addSss ss)),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   218
      prune_params_tac]
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   219
  end;
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   220
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   221
fun auto_tac css = mk_auto_tac css 4 2;
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   222
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   223
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   224
(* force_tac *)
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   225
4659
a78ecc7341e3 added smart_tac
oheimb
parents: 4652
diff changeset
   226
(* aimed to solve the given subgoal totally, using whatever tools possible *)
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   227
fun force_tac (cs, ss) =
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   228
  let val cs' = cs addss ss in
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   229
    SELECT_GOAL (EVERY [
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   230
      Classical.clarify_tac cs' 1,
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   231
      IF_UNSOLVED (Simplifier.asm_full_simp_tac ss 1),
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   232
      ALLGOALS (Classical.first_best_tac cs')])
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   233
  end;
4659
a78ecc7341e3 added smart_tac
oheimb
parents: 4652
diff changeset
   234
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 4888
diff changeset
   235
9805
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   236
(* basic combinations *)
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   237
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   238
fun ADDSS tac (cs, ss) = let val cs' = cs addss ss in tac cs' end;
9591
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   239
9805
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   240
val fast_simp_tac = ADDSS Classical.fast_tac;
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   241
val slow_simp_tac = ADDSS Classical.slow_tac;
10b617bdd028 added "slowsimp", "bestsimp";
wenzelm
parents: 9772
diff changeset
   242
val best_simp_tac = ADDSS Classical.best_tac;
9591
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   243
590d36e059d1 added "fastsimp";
wenzelm
parents: 9506
diff changeset
   244
8639
31bcb6b64d60 added change_global/local_css;
wenzelm
parents: 8469
diff changeset
   245
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42373
diff changeset
   246
(** access clasimpset **)
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   247
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   248
(* attributes *)
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   249
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   250
fun attrib f = Thm.declaration_attribute (fn th => map_css (fn css => f (css, [th])));
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   251
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
   252
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   253
val iff_add = attrib (op addIffs);
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   254
val iff_add' = attrib' addIffs_generic;
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   255
val iff_del = attrib (op delIffs) o attrib' delIffs_generic;
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   256
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30513
diff changeset
   257
fun iff_att x = (Scan.lift
18688
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   258
 (Args.del >> K iff_del ||
abf0f018b5ec generic attributes;
wenzelm
parents: 18630
diff changeset
   259
  Scan.option Args.add -- Args.query >> K iff_add' ||
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30513
diff changeset
   260
  Scan.option Args.add >> K iff_add)) x;
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   261
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   262
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   263
(* method modifiers *)
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   264
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   265
val iffN = "iff";
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   266
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   267
val iff_modifiers =
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   268
 [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add): Method.modifier),
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   269
  Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add'),
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   270
  Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del)];
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   271
8469
482c301041b4 include Splitter.split_modifiers;
wenzelm
parents: 8168
diff changeset
   272
val clasimp_modifiers =
9860
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   273
  Simplifier.simp_modifiers @ Splitter.split_modifiers @
5c5efed691b9 added 'iff' declarations;
wenzelm
parents: 9805
diff changeset
   274
  Classical.cla_modifiers @ iff_modifiers;
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
(* methods *)
5926
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   278
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   279
fun clasimp_meth tac ctxt = METHOD (fn facts =>
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   280
  ALLGOALS (Method.insert_tac facts) THEN tac (clasimpset_of ctxt));
7132
5c31d06ead04 eliminated METHOD0 in favour of same_tac;
wenzelm
parents: 5985
diff changeset
   281
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   282
fun clasimp_meth' tac ctxt = METHOD (fn facts =>
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   283
  HEADGOAL (Method.insert_tac facts THEN' tac (clasimpset_of ctxt)));
5926
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   284
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   285
30541
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   286
fun clasimp_method' tac =
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   287
  Method.sections clasimp_modifiers >> K (clasimp_meth' tac);
9772
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   288
30541
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   289
val auto_method =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36601
diff changeset
   290
  Scan.lift (Scan.option (Parse.nat -- Parse.nat)) --|
35613
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   291
    Method.sections clasimp_modifiers >>
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   292
  (fn NONE => clasimp_meth (CHANGED_PROP o auto_tac)
9d3ff36ad4e1 eliminated Args.bang_facts (legacy feature);
wenzelm
parents: 33369
diff changeset
   293
    | 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
   294
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   295
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   296
(* theory setup *)
c07777210a69 auto method: opt args;
wenzelm
parents: 9591
diff changeset
   297
26497
1873915c64a9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26425
diff changeset
   298
val clasimp_setup =
30541
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   299
  Attrib.setup @{binding iff} iff_att "declaration of Simplifier / Classical rules" #>
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   300
  Method.setup @{binding fastsimp} (clasimp_method' fast_simp_tac) "combined fast and simp" #>
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   301
  Method.setup @{binding slowsimp} (clasimp_method' slow_simp_tac) "combined slow and simp" #>
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   302
  Method.setup @{binding bestsimp} (clasimp_method' best_simp_tac) "combined best and simp" #>
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   303
  Method.setup @{binding force} (clasimp_method' force_tac) "force" #>
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   304
  Method.setup @{binding auto} auto_method "auto" #>
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   305
  Method.setup @{binding clarsimp} (clasimp_method' (CHANGED_PROP oo clarsimp_tac))
9f168bdc468a simplified method setup;
wenzelm
parents: 30528
diff changeset
   306
    "clarify simplified goal";
5926
58f9ca06b76b method setup;
wenzelm
parents: 5756
diff changeset
   307
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents:
diff changeset
   308
end;