src/FOL/simpdata.ML
author wenzelm
Wed, 22 Apr 2020 19:22:43 +0200
changeset 71787 acfe72ff00c2
parent 69593 3dda49e08b9d
child 71916 435cdc772110
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9889
8802b140334c rulify setup;
wenzelm
parents: 9851
diff changeset
     1
(*  Title:      FOL/simpdata.ML
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
     3
    Copyright   1994  University of Cambridge
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
9889
8802b140334c rulify setup;
wenzelm
parents: 9851
diff changeset
     5
Simplification data for FOL.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
     8
(*Make meta-equalities.  The operator below is Trueprop*)
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
     9
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    10
fun mk_meta_eq th =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    11
  (case Thm.concl_of th of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    12
    _ $ (Const(\<^const_name>\<open>eq\<close>,_)$_$_)   => th RS @{thm eq_reflection}
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    13
  | _ $ (Const(\<^const_name>\<open>iff\<close>,_)$_$_) => th RS @{thm iff_reflection}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    14
  | _ => error "conclusion must be a =-equality or <->");
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    15
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    16
fun mk_eq th =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    17
  (case Thm.concl_of th of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    18
    Const(\<^const_name>\<open>Pure.eq\<close>,_)$_$_ => th
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    19
  | _ $ (Const(\<^const_name>\<open>eq\<close>,_)$_$_)   => mk_meta_eq th
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    20
  | _ $ (Const(\<^const_name>\<open>iff\<close>,_)$_$_) => mk_meta_eq th
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    21
  | _ $ (Const(\<^const_name>\<open>Not\<close>,_)$_)      => th RS @{thm iff_reflection_F}
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
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*)
36546
a9873318fe30 proper context for rule_by_tactic;
wenzelm
parents: 36543
diff changeset
    25
fun mk_meta_prems ctxt =
a9873318fe30 proper context for rule_by_tactic;
wenzelm
parents: 36543
diff changeset
    26
    rule_by_tactic ctxt
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
    27
      (REPEAT_FIRST (resolve_tac ctxt [@{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 ==)*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45659
diff changeset
    30
fun mk_meta_cong ctxt rl =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45659
diff changeset
    31
  Drule.zero_var_indexes (mk_meta_eq (mk_meta_prems ctxt rl))
35021
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 32957
diff changeset
    32
    handle THM _ =>
c839a4c670c6 renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents: 32957
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 =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    36
  [(\<^const_name>\<open>imp\<close>, [@{thm mp}]), (\<^const_name>\<open>conj\<close>, [@{thm conjunct1}, @{thm conjunct2}]),
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    37
   (\<^const_name>\<open>All\<close>, [@{thm spec}]), (\<^const_name>\<open>True\<close>, []), (\<^const_name>\<open>False\<close>, [])];
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    38
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    39
fun mk_atomize pairs =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    40
  let
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    41
    fun atoms th =
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    42
      (case Thm.concl_of th of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    43
         Const(\<^const_name>\<open>Trueprop\<close>,_) $ p =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    44
           (case head_of p of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    45
              Const(a,_) =>
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    46
                (case AList.lookup (op =) pairs a of
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    47
                   SOME(rls) => maps atoms ([th] RL rls)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    48
                 | NONE => [th])
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    49
            | _ => [th])
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    50
       | _ => [th])
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    51
  in atoms end;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
    52
60822
4f58f3662e7d more explicit context;
wenzelm
parents: 59970
diff changeset
    53
fun mksimps pairs ctxt = map mk_eq o mk_atomize pairs o Variable.gen_all ctxt;
981
864370666a24 Defined addss to perform simplification in a claset.
lcp
parents: 942
diff changeset
    54
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
    55
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    56
(** make simplification procedures for quantifier elimination **)
42458
5dfae6d348fd misc tuning;
wenzelm
parents: 42455
diff changeset
    57
structure Quantifier1 = Quantifier1
5dfae6d348fd misc tuning;
wenzelm
parents: 42455
diff changeset
    58
(
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    59
  (*abstract syntax*)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    60
  fun dest_eq (Const (\<^const_name>\<open>eq\<close>, _) $ s $ t) = SOME (s, t)
42460
1805c67dc7aa simplified Data signature;
wenzelm
parents: 42458
diff changeset
    61
    | dest_eq _ = NONE
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    62
  fun dest_conj (Const (\<^const_name>\<open>conj\<close>, _) $ s $ t) = SOME (s, t)
42460
1805c67dc7aa simplified Data signature;
wenzelm
parents: 42458
diff changeset
    63
    | dest_conj _ = NONE
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    64
  fun dest_imp (Const (\<^const_name>\<open>imp\<close>, _) $ s $ t) = SOME (s, t)
42460
1805c67dc7aa simplified Data signature;
wenzelm
parents: 42458
diff changeset
    65
    | dest_imp _ = NONE
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    66
  val conj = FOLogic.conj
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    67
  val imp  = FOLogic.imp
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    68
  (*rules*)
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    69
  val iff_reflection = @{thm iff_reflection}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    70
  val iffI = @{thm iffI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    71
  val iff_trans = @{thm iff_trans}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    72
  val conjI= @{thm conjI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    73
  val conjE= @{thm conjE}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    74
  val impI = @{thm impI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    75
  val mp   = @{thm mp}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    76
  val uncurry = @{thm uncurry}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    77
  val exI  = @{thm exI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    78
  val exE  = @{thm exE}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    79
  val iff_allI = @{thm iff_allI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    80
  val iff_exI = @{thm iff_exI}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    81
  val all_comm = @{thm all_comm}
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    82
  val ex_comm = @{thm ex_comm}
42458
5dfae6d348fd misc tuning;
wenzelm
parents: 42455
diff changeset
    83
);
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    84
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    85
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    86
(*** Case splitting ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
32177
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    88
structure Splitter = Splitter
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    89
(
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
    90
  val context = \<^context>
32177
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    91
  val mk_eq = mk_eq
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
    92
  val meta_eq_to_iff = @{thm meta_eq_to_iff}
32177
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    93
  val iffD = @{thm iffD2}
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    94
  val disjE = @{thm disjE}
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    95
  val conjE = @{thm conjE}
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    96
  val exE = @{thm exE}
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    97
  val contrapos = @{thm contrapos}
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    98
  val contrapos2 = @{thm contrapos2}
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
    99
  val notnotD = @{thm notnotD}
63637
9a57baa15e1b adapted ZF,FOL,CCL,LCF to modified splitter
nipkow
parents: 60822
diff changeset
   100
  val safe_tac = Cla.safe_tac
32177
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
   101
);
1722
bb326972ede6 Added split_inside_tac.
berghofe
parents: 1459
diff changeset
   102
32177
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
   103
val split_tac = Splitter.split_tac;
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   104
val split_inside_tac = Splitter.split_inside_tac;
32177
bc02c5bfcb5b renamed functor SplitterFun to Splitter, require explicit theory;
wenzelm
parents: 32155
diff changeset
   105
val split_asm_tac = Splitter.split_asm_tac;
4325
e72cba5af6c5 addsplits now in FOL, ZF too
paulson
parents: 4203
diff changeset
   106
e72cba5af6c5 addsplits now in FOL, ZF too
paulson
parents: 4203
diff changeset
   107
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   108
(*** Standard simpsets ***)
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   109
26288
89b9f7c18631 eliminated out-of-scope proofs (cf. theory IFOL and FOL);
wenzelm
parents: 22822
diff changeset
   110
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
   111
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45659
diff changeset
   112
fun unsafe_solver ctxt =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   113
  FIRST' [resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt),
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58957
diff changeset
   114
    assume_tac ctxt,
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   115
    eresolve_tac ctxt @{thms FalseE}];
43597
b4a093e755db tuned signature;
wenzelm
parents: 43596
diff changeset
   116
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   117
(*No premature instantiation of variables during simplification*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45659
diff changeset
   118
fun safe_solver ctxt =
58957
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58838
diff changeset
   119
  FIRST' [match_tac ctxt (triv_rls @ Simplifier.prems_of ctxt),
c9e744ea8a38 proper context for match_tac etc.;
wenzelm
parents: 58838
diff changeset
   120
    eq_assume_tac, ematch_tac ctxt @{thms FalseE}];
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   121
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   122
(*No simprules, but basic infastructure for simplification*)
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17875
diff changeset
   123
val FOL_basic_ss =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 63637
diff changeset
   124
  empty_simpset \<^context>
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   125
  setSSolver (mk_solver "FOL safe" safe_solver)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   126
  setSolver (mk_solver "FOL unsafe" unsafe_solver)
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45620
diff changeset
   127
  |> Simplifier.set_subgoaler asm_simp_tac
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 45620
diff changeset
   128
  |> Simplifier.set_mksimps (mksimps mksimps_pairs)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45659
diff changeset
   129
  |> Simplifier.set_mkcong mk_meta_cong
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45659
diff changeset
   130
  |> simpset_of;
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   131
54998
8601434fa334 tuned signature;
wenzelm
parents: 51717
diff changeset
   132
fun unfold_tac ctxt ths =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45659
diff changeset
   133
  ALLGOALS (full_simp_tac (clear_simpset (put_simpset FOL_basic_ss ctxt) addsimps ths));
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 16019
diff changeset
   134
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   135
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 5115
diff changeset
   136
(*** integration of simplifier with classical reasoner ***)
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   137
42478
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   138
structure Clasimp = Clasimp
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   139
(
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   140
  structure Simplifier = Simplifier
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   141
    and Splitter = Splitter
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   142
    and Classical = Cla
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   143
    and Blast = Blast
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   144
  val iffD1 = @{thm iffD1}
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   145
  val iffD2 = @{thm iffD2}
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   146
  val notE = @{thm notE}
8a526c010c3b modernized Clasimp setup;
wenzelm
parents: 42460
diff changeset
   147
);
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents: 4649
diff changeset
   148
open Clasimp;
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   149