src/FOL/simpdata.ML
author wenzelm
Wed, 17 Sep 2008 21:27:08 +0200
changeset 28262 aa7ca36d67fd
parent 27338 2cd6c60cc10b
child 30609 983e8b6e4e69
permissions -rw-r--r--
back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9889
8802b140334c rulify setup;
wenzelm
parents: 9851
diff changeset
     1
(*  Title:      FOL/simpdata.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
     4
    Copyright   1994  University of Cambridge
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
9889
8802b140334c rulify setup;
wenzelm
parents: 9851
diff changeset
     6
Simplification data for FOL.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
     9
(*Make meta-equalities.  The operator below is Trueprop*)
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    10
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    11
fun mk_meta_eq th = case concl_of th of
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    12
    _ $ (Const("op =",_)$_$_)   => th RS @{thm eq_reflection}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    13
  | _ $ (Const("op <->",_)$_$_) => th RS @{thm iff_reflection}
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    14
  | _                           =>
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    15
  error("conclusion must be a =-equality or <->");;
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    16
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    17
fun mk_eq th = case concl_of th of
394
432bb9995893 Modified mk_meta_eq to leave meta-equlities on unchanged.
nipkow
parents: 371
diff changeset
    18
    Const("==",_)$_$_           => th
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    19
  | _ $ (Const("op =",_)$_$_)   => mk_meta_eq th
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    20
  | _ $ (Const("op <->",_)$_$_) => mk_meta_eq th
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    21
  | _ $ (Const("Not",_)$_)      => th RS @{thm iff_reflection_F}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    22
  | _                           => th RS @{thm iff_reflection_T};
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
6114
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    24
(*Replace premises x=y, X<->Y by X==Y*)
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    25
val mk_meta_prems =
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    26
    rule_by_tactic
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    27
      (REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
6114
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    28
9713
2c5b42311eb0 cong setup now part of Simplifier;
wenzelm
parents: 9300
diff changeset
    29
(*Congruence rules for = or <-> (instead of ==)*)
6114
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    30
fun mk_meta_cong rl =
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    31
  standard(mk_meta_eq (mk_meta_prems rl))
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    32
  handle THM _ =>
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    33
  error("Premises and conclusion of congruence rules must use =-equality or <->");
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    34
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    35
val mksimps_pairs =
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    36
  [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]),
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    37
   ("All", [@{thm spec}]), ("True", []), ("False", [])];
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    38
16019
0e1405402d53 Simplifier already setup in Pure;
wenzelm
parents: 15570
diff changeset
    39
(* ###FIXME: move to simplifier.ML
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    40
val mk_atomize:      (string * thm list) list -> thm -> thm list
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    41
*)
16019
0e1405402d53 Simplifier already setup in Pure;
wenzelm
parents: 15570
diff changeset
    42
(* ###FIXME: move to simplifier.ML *)
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    43
fun mk_atomize pairs =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    44
  let fun atoms th =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    45
        (case concl_of th of
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    46
           Const("Trueprop",_) $ p =>
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    47
             (case head_of p of
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    48
                Const(a,_) =>
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 17002
diff changeset
    49
                  (case AList.lookup (op =) pairs a of
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    50
                     SOME(rls) => List.concat (map atoms ([th] RL rls))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
    51
                   | NONE => [th])
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    52
              | _ => [th])
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    53
         | _ => [th])
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    54
  in atoms end;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    55
12725
7ede865e1fe5 renamed forall_elim_vars_safe to gen_all;
wenzelm
parents: 12720
diff changeset
    56
fun mksimps pairs = (map mk_eq o mk_atomize pairs o gen_all);
981
864370666a24 Defined addss to perform simplification in a claset.
lcp
parents: 942
diff changeset
    57
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
    58
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    59
(** make simplification procedures for quantifier elimination **)
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    60
structure Quantifier1 = Quantifier1Fun(
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    61
struct
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    62
  (*abstract syntax*)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
    63
  fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
    64
    | dest_eq _ = NONE;
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
    65
  fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
    66
    | dest_conj _ = NONE;
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
    67
  fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
    68
    | dest_imp _ = NONE;
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    69
  val conj = FOLogic.conj
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    70
  val imp  = FOLogic.imp
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    71
  (*rules*)
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    72
  val iff_reflection = @{thm iff_reflection}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    73
  val iffI = @{thm iffI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    74
  val iff_trans = @{thm iff_trans}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    75
  val conjI= @{thm conjI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    76
  val conjE= @{thm conjE}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    77
  val impI = @{thm impI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    78
  val mp   = @{thm mp}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    79
  val uncurry = @{thm uncurry}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    80
  val exI  = @{thm exI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    81
  val exE  = @{thm exE}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    82
  val iff_allI = @{thm iff_allI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    83
  val iff_exI = @{thm iff_exI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    84
  val all_comm = @{thm all_comm}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    85
  val ex_comm = @{thm ex_comm}
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    86
end);
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    87
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
    88
val defEX_regroup =
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16019
diff changeset
    89
  Simplifier.simproc (the_context ())
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
    90
    "defined EX" ["EX x. P(x)"] Quantifier1.rearrange_ex;
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    91
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    92
val defALL_regroup =
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16019
diff changeset
    93
  Simplifier.simproc (the_context ())
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
    94
    "defined ALL" ["ALL x. P(x)"] Quantifier1.rearrange_all;
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    95
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    96
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    97
(*** Case splitting ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    99
structure SplitterData =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   100
  struct
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   101
  structure Simplifier = Simplifier
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
   102
  val mk_eq          = mk_eq
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   103
  val meta_eq_to_iff = @{thm meta_eq_to_iff}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   104
  val iffD           = @{thm iffD2}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   105
  val disjE          = @{thm disjE}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   106
  val conjE          = @{thm conjE}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   107
  val exE            = @{thm exE}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   108
  val contrapos      = @{thm contrapos}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   109
  val contrapos2     = @{thm contrapos2}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   110
  val notnotD        = @{thm notnotD}
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   111
  end;
1722
bb326972ede6 Added split_inside_tac.
berghofe
parents: 1459
diff changeset
   112
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   113
structure Splitter = SplitterFun(SplitterData);
1722
bb326972ede6 Added split_inside_tac.
berghofe
parents: 1459
diff changeset
   114
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   115
val split_tac        = Splitter.split_tac;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   116
val split_inside_tac = Splitter.split_inside_tac;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   117
val split_asm_tac    = Splitter.split_asm_tac;
5307
6a699d5cdef4 minor adaption for SML/NJ
oheimb
parents: 5304
diff changeset
   118
val op addsplits     = Splitter.addsplits;
6a699d5cdef4 minor adaption for SML/NJ
oheimb
parents: 5304
diff changeset
   119
val op delsplits     = Splitter.delsplits;
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   120
val Addsplits        = Splitter.Addsplits;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   121
val Delsplits        = Splitter.Delsplits;
4325
e72cba5af6c5 addsplits now in FOL, ZF too
paulson
parents: 4203
diff changeset
   122
e72cba5af6c5 addsplits now in FOL, ZF too
paulson
parents: 4203
diff changeset
   123
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   124
(*** Standard simpsets ***)
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   125
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   126
val triv_rls = [@{thm TrueI}, @{thm refl}, reflexive_thm, @{thm iff_refl}, @{thm notFalseI}];
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   127
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   128
fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls @ prems),
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   129
                                 atac, etac @{thm FalseE}];
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   130
(*No premature instantiation of variables during simplification*)
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   131
fun   safe_solver prems = FIRST'[match_tac (triv_rls @ prems),
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   132
                                 eq_assume_tac, ematch_tac [@{thm FalseE}]];
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   133
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   134
(*No simprules, but basic infastructure for simplification*)
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17875
diff changeset
   135
val FOL_basic_ss =
28262
aa7ca36d67fd back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
wenzelm
parents: 27338
diff changeset
   136
  Simplifier.theory_context (the_context ()) empty_ss
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   137
  setsubgoaler asm_simp_tac
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   138
  setSSolver (mk_solver "FOL safe" safe_solver)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   139
  setSolver (mk_solver "FOL unsafe" unsafe_solver)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   140
  setmksimps (mksimps mksimps_pairs)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   141
  setmkcong mk_meta_cong;
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   142
18324
d1c4b1112e33 unfold_tac: static evaluation of simpset;
wenzelm
parents: 17892
diff changeset
   143
fun unfold_tac ths =
d1c4b1112e33 unfold_tac: static evaluation of simpset;
wenzelm
parents: 17892
diff changeset
   144
  let val ss0 = Simplifier.clear_ss FOL_basic_ss addsimps ths
d1c4b1112e33 unfold_tac: static evaluation of simpset;
wenzelm
parents: 17892
diff changeset
   145
  in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end;
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16019
diff changeset
   146
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   147
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   148
(*intuitionistic simprules only*)
21539
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 20223
diff changeset
   149
val IFOL_ss =
c5cf9243ad62 converted legacy ML scripts;
wenzelm
parents: 20223
diff changeset
   150
  FOL_basic_ss
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   151
  addsimps (@{thms meta_simps} @ @{thms IFOL_simps} @ @{thms int_ex_simps} @ @{thms int_all_simps})
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   152
  addsimprocs [defALL_regroup, defEX_regroup]    
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   153
  addcongs [@{thm imp_cong}];
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   154
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   155
(*classical simprules too*)
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   156
val FOL_ss = IFOL_ss addsimps (@{thms cla_simps} @ @{thms cla_ex_simps} @ @{thms cla_all_simps});
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   157
26496
49ae9456eba9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 26288
diff changeset
   158
val simpsetup = Simplifier.map_simpset (K FOL_ss);
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   159
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   160
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 5115
diff changeset
   161
(*** integration of simplifier with classical reasoner ***)
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   162
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 5115
diff changeset
   163
structure Clasimp = ClasimpFun
8472
50a653f8b8ea clasimp: include Splitter;
wenzelm
parents: 7570
diff changeset
   164
 (structure Simplifier = Simplifier and Splitter = Splitter
9851
e22db9397e17 iff declarations moved to clasimp.ML;
wenzelm
parents: 9713
diff changeset
   165
  and Classical  = Cla and Blast = Blast
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   166
  val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE});
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents: 4649
diff changeset
   167
open Clasimp;
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   168
27338
2cd6c60cc10b ML_Antiquote.value;
wenzelm
parents: 26496
diff changeset
   169
ML_Antiquote.value "clasimpset"
2cd6c60cc10b ML_Antiquote.value;
wenzelm
parents: 26496
diff changeset
   170
  (Scan.succeed "Clasimp.local_clasimpset_of (ML_Context.the_local_context ())");
22128
cdd92316dd31 added @{clasimpset};
wenzelm
parents: 21539
diff changeset
   171
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   172
val FOL_css = (FOL_cs, FOL_ss);