src/FOL/simpdata.ML
author huffman
Thu, 26 May 2005 02:23:27 +0200
changeset 16081 81a4b4a245b0
parent 16019 0e1405402d53
child 17002 fb9261990ffe
permissions -rw-r--r--
cleaned up, added cpair_less and cpair_eq_pair, removed some obsolete stuff
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
9300
ee5c9672d208 AddIffs now available for FOL, ZF
paulson
parents: 8643
diff changeset
     9
5496
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
    10
(* Elimination of True from asumptions: *)
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
    11
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    12
bind_thm ("True_implies_equals", prove_goal IFOL.thy
5496
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
    13
 "(True ==> PROP P) == PROP P"
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
    14
(K [rtac equal_intr_rule 1, atac 2,
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
    15
          METAHYPS (fn prems => resolve_tac prems 1) 1,
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    16
          rtac TrueI 1]));
5496
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
    17
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
    18
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
(*** Rewrite rules ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    21
fun int_prove_fun s =
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    22
 (writeln s;
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    23
  prove_goal IFOL.thy s
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    24
   (fn prems => [ (cut_facts_tac prems 1),
2601
b301958c465d Renamed structure Int (intuitionistic prover) to IntPr to prevent clash
paulson
parents: 2469
diff changeset
    25
                  (IntPr.fast_tac 1) ]));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    27
bind_thms ("conj_simps", map int_prove_fun
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
    28
 ["P & True <-> P",      "True & P <-> P",
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
  "P & False <-> False", "False & P <-> False",
2801
56948cb1a1f9 Added P&P&Q <-> P&Q and P|P|Q <-> P|Q
nipkow
parents: 2727
diff changeset
    30
  "P & P <-> P", "P & P & Q <-> P & Q",
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
    31
  "P & ~P <-> False",    "~P & P <-> False",
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    32
  "(P & Q) & R <-> P & (Q & R)"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    34
bind_thms ("disj_simps", map int_prove_fun
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
    35
 ["P | True <-> True",  "True | P <-> True",
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
    36
  "P | False <-> P",    "False | P <-> P",
2801
56948cb1a1f9 Added P&P&Q <-> P&Q and P|P|Q <-> P|Q
nipkow
parents: 2727
diff changeset
    37
  "P | P <-> P", "P | P | Q <-> P | Q",
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    38
  "(P | Q) | R <-> P | (Q | R)"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    40
bind_thms ("not_simps", map int_prove_fun
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    41
 ["~(P|Q)  <-> ~P & ~Q",
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    42
  "~ False <-> True",   "~ True <-> False"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    44
bind_thms ("imp_simps", map int_prove_fun
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
    45
 ["(P --> False) <-> ~P",       "(P --> True) <-> True",
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    46
  "(False --> P) <-> True",     "(True --> P) <-> P",
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    47
  "(P --> P) <-> True",         "(P --> ~P) <-> ~P"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    49
bind_thms ("iff_simps", map int_prove_fun
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
    50
 ["(True <-> P) <-> P",         "(P <-> True) <-> P",
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
  "(P <-> P) <-> True",
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    52
  "(False <-> P) <-> ~P",       "(P <-> False) <-> ~P"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    53
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    54
(*The x=t versions are needed for the simplification procedures*)
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    55
bind_thms ("quant_simps", map int_prove_fun
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    56
 ["(ALL x. P) <-> P",
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    57
  "(ALL x. x=t --> P(x)) <-> P(t)",
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    58
  "(ALL x. t=x --> P(x)) <-> P(t)",
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
    59
  "(EX x. P) <-> P",
13149
773657d466cb better simplification of trivial existential equalities
paulson
parents: 12825
diff changeset
    60
  "EX x. x=t", "EX x. t=x",
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    61
  "(EX x. x=t & P(x)) <-> P(t)",
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    62
  "(EX x. t=x & P(x)) <-> P(t)"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
(*These are NOT supplied by default!*)
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    65
bind_thms ("distrib_simps", map int_prove_fun
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    66
 ["P & (Q | R) <-> P&Q | P&R",
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    67
  "(Q | R) & P <-> Q&P | R&P",
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    68
  "(P | Q --> R) <-> (P --> R) & (Q --> R)"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    70
(** Conversion into rewrite rules **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    72
bind_thm ("P_iff_F", int_prove_fun "~P ==> (P <-> False)");
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    73
bind_thm ("iff_reflection_F", P_iff_F RS iff_reflection);
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    74
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    75
bind_thm ("P_iff_T", int_prove_fun "P ==> (P <-> True)");
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
    76
bind_thm ("iff_reflection_T", P_iff_T RS iff_reflection);
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    77
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    78
(*Make meta-equalities.  The operator below is Trueprop*)
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    79
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    80
fun mk_meta_eq th = case concl_of th of
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    81
    _ $ (Const("op =",_)$_$_)   => th RS eq_reflection
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    82
  | _ $ (Const("op <->",_)$_$_) => th RS iff_reflection
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    83
  | _                           =>
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    84
  error("conclusion must be a =-equality or <->");;
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    85
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    86
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
    87
    Const("==",_)$_$_           => th
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    88
  | _ $ (Const("op =",_)$_$_)   => mk_meta_eq th
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
    89
  | _ $ (Const("op <->",_)$_$_) => mk_meta_eq th
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    90
  | _ $ (Const("Not",_)$_)      => th RS iff_reflection_F
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
    91
  | _                           => th RS iff_reflection_T;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    92
6114
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    93
(*Replace premises x=y, X<->Y by X==Y*)
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    94
val mk_meta_prems =
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
    95
    rule_by_tactic
6114
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    96
      (REPEAT_FIRST (resolve_tac [meta_eq_to_obj_eq, def_imp_iff]));
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    97
9713
2c5b42311eb0 cong setup now part of Simplifier;
wenzelm
parents: 9300
diff changeset
    98
(*Congruence rules for = or <-> (instead of ==)*)
6114
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
    99
fun mk_meta_cong rl =
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
   100
  standard(mk_meta_eq (mk_meta_prems rl))
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
   101
  handle THM _ =>
45958e54d72e congruence rules finally use == instead of = and <->
paulson
parents: 5555
diff changeset
   102
  error("Premises and conclusion of congruence rules must use =-equality or <->");
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
   103
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   104
val mksimps_pairs =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   105
  [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   106
   ("All", [spec]), ("True", []), ("False", [])];
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   107
16019
0e1405402d53 Simplifier already setup in Pure;
wenzelm
parents: 15570
diff changeset
   108
(* ###FIXME: move to simplifier.ML
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   109
val mk_atomize:      (string * thm list) list -> thm -> thm list
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   110
*)
16019
0e1405402d53 Simplifier already setup in Pure;
wenzelm
parents: 15570
diff changeset
   111
(* ###FIXME: move to simplifier.ML *)
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   112
fun mk_atomize pairs =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   113
  let fun atoms th =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   114
        (case concl_of th of
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   115
           Const("Trueprop",_) $ p =>
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   116
             (case head_of p of
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   117
                Const(a,_) =>
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   118
                  (case assoc(pairs,a) of
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   119
                     SOME(rls) => List.concat (map atoms ([th] RL rls))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
   120
                   | NONE => [th])
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   121
              | _ => [th])
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   122
         | _ => [th])
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   123
  in atoms end;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   124
12725
7ede865e1fe5 renamed forall_elim_vars_safe to gen_all;
wenzelm
parents: 12720
diff changeset
   125
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
   126
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   127
(*** Classical laws ***)
282
731b27c90d2f FOL/simpdata: tidied
lcp
parents: 215
diff changeset
   128
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   129
fun prove_fun s =
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   130
 (writeln s;
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6391
diff changeset
   131
  prove_goal (the_context ()) s
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   132
   (fn prems => [ (cut_facts_tac prems 1),
1459
d12da312eff4 expanded tabs
clasohm
parents: 1088
diff changeset
   133
                  (Cla.fast_tac FOL_cs 1) ]));
745
45a789407806 added blank line
lcp
parents: 429
diff changeset
   134
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   135
(*Avoids duplication of subgoals after expand_if, when the true and false
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   136
  cases boil down to the same thing.*)
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   137
bind_thm ("cases_simp", prove_fun "(P --> Q) & (~P --> Q) <-> Q");
1953
832ccc1dba95 Introduction of miniscoping for FOL
paulson
parents: 1914
diff changeset
   138
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   139
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   140
(*** Miniscoping: pushing quantifiers in
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   141
     We do NOT distribute of ALL over &, or dually that of EX over |
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   142
     Baaz and Leitsch, On Skolemization and Proof Complexity (1994)
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   143
     show that this step can increase proof length!
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   144
***)
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   145
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   146
(*existential miniscoping*)
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   147
bind_thms ("int_ex_simps", map int_prove_fun
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   148
 ["(EX x. P(x) & Q) <-> (EX x. P(x)) & Q",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   149
  "(EX x. P & Q(x)) <-> P & (EX x. Q(x))",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   150
  "(EX x. P(x) | Q) <-> (EX x. P(x)) | Q",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   151
  "(EX x. P | Q(x)) <-> P | (EX x. Q(x))"]);
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   152
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   153
(*classical rules*)
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   154
bind_thms ("cla_ex_simps", map prove_fun
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   155
 ["(EX x. P(x) --> Q) <-> (ALL x. P(x)) --> Q",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   156
  "(EX x. P --> Q(x)) <-> P --> (EX x. Q(x))"]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   157
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   158
bind_thms ("ex_simps", int_ex_simps @ cla_ex_simps);
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   159
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   160
(*universal miniscoping*)
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   161
bind_thms ("int_all_simps", map int_prove_fun
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   162
 ["(ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   163
  "(ALL x. P & Q(x)) <-> P & (ALL x. Q(x))",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   164
  "(ALL x. P(x) --> Q) <-> (EX x. P(x)) --> Q",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   165
  "(ALL x. P --> Q(x)) <-> P --> (ALL x. Q(x))"]);
1953
832ccc1dba95 Introduction of miniscoping for FOL
paulson
parents: 1914
diff changeset
   166
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   167
(*classical rules*)
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   168
bind_thms ("cla_all_simps", map prove_fun
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   169
 ["(ALL x. P(x) | Q) <-> (ALL x. P(x)) | Q",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   170
  "(ALL x. P | Q(x)) <-> P | (ALL x. Q(x))"]);
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   171
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   172
bind_thms ("all_simps", int_all_simps @ cla_all_simps);
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   173
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   174
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   175
(*** Named rewrite rules proved for IFOL ***)
1953
832ccc1dba95 Introduction of miniscoping for FOL
paulson
parents: 1914
diff changeset
   176
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   177
fun int_prove nm thm  = qed_goal nm IFOL.thy thm
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   178
    (fn prems => [ (cut_facts_tac prems 1),
2601
b301958c465d Renamed structure Int (intuitionistic prover) to IntPr to prevent clash
paulson
parents: 2469
diff changeset
   179
                   (IntPr.fast_tac 1) ]);
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   180
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6391
diff changeset
   181
fun prove nm thm  = qed_goal nm (the_context ()) thm (fn _ => [Blast_tac 1]);
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   182
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   183
int_prove "conj_commute" "P&Q <-> Q&P";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   184
int_prove "conj_left_commute" "P&(Q&R) <-> Q&(P&R)";
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   185
bind_thms ("conj_comms", [conj_commute, conj_left_commute]);
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   186
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   187
int_prove "disj_commute" "P|Q <-> Q|P";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   188
int_prove "disj_left_commute" "P|(Q|R) <-> Q|(P|R)";
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   189
bind_thms ("disj_comms", [disj_commute, disj_left_commute]);
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   190
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   191
int_prove "conj_disj_distribL" "P&(Q|R) <-> (P&Q | P&R)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   192
int_prove "conj_disj_distribR" "(P|Q)&R <-> (P&R | Q&R)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   193
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   194
int_prove "disj_conj_distribL" "P|(Q&R) <-> (P|Q) & (P|R)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   195
int_prove "disj_conj_distribR" "(P&Q)|R <-> (P|R) & (Q|R)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   196
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   197
int_prove "imp_conj_distrib" "(P --> (Q&R)) <-> (P-->Q) & (P-->R)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   198
int_prove "imp_conj"         "((P&Q)-->R)   <-> (P --> (Q --> R))";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   199
int_prove "imp_disj"         "(P|Q --> R)   <-> (P-->R) & (Q-->R)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   200
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   201
prove "imp_disj1" "(P-->Q) | R <-> (P-->Q | R)";
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   202
prove "imp_disj2" "Q | (P-->R) <-> (P-->Q | R)";
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   203
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   204
int_prove "de_Morgan_disj" "(~(P | Q)) <-> (~P & ~Q)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   205
prove     "de_Morgan_conj" "(~(P & Q)) <-> (~P | ~Q)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   206
12765
fb3f9887d0b7 new theorem
paulson
parents: 12725
diff changeset
   207
prove     "not_imp" "~(P --> Q) <-> (P & ~Q)";
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   208
prove     "not_iff" "~(P <-> Q) <-> (P <-> ~Q)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   209
3835
9a5a4e123859 fixed dots;
wenzelm
parents: 3610
diff changeset
   210
prove     "not_all" "(~ (ALL x. P(x))) <-> (EX x.~P(x))";
9a5a4e123859 fixed dots;
wenzelm
parents: 3610
diff changeset
   211
prove     "imp_all" "((ALL x. P(x)) --> Q) <-> (EX x. P(x) --> Q)";
9a5a4e123859 fixed dots;
wenzelm
parents: 3610
diff changeset
   212
int_prove "not_ex"  "(~ (EX x. P(x))) <-> (ALL x.~P(x))";
1914
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   213
int_prove "imp_ex" "((EX x. P(x)) --> Q) <-> (ALL x. P(x) --> Q)";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   214
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   215
int_prove "ex_disj_distrib"
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   216
    "(EX x. P(x) | Q(x)) <-> ((EX x. P(x)) | (EX x. Q(x)))";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   217
int_prove "all_conj_distrib"
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   218
    "(ALL x. P(x) & Q(x)) <-> ((ALL x. P(x)) & (ALL x. Q(x)))";
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   219
86b095835de9 Added a lot of basic laws, from HOL/simpdata
paulson
parents: 1722
diff changeset
   220
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   221
local
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   222
val uncurry = prove_goal (the_context()) "P --> Q --> R ==> P & Q --> R"
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   223
              (fn prems => [cut_facts_tac prems 1, Blast_tac 1]);
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   224
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   225
val iff_allI = allI RS
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   226
    prove_goal (the_context()) "ALL x. P(x) <-> Q(x) ==> (ALL x. P(x)) <-> (ALL x. Q(x))"
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   227
               (fn prems => [cut_facts_tac prems 1, Blast_tac 1])
12526
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   228
val iff_exI = allI RS
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   229
    prove_goal (the_context()) "ALL x. P(x) <-> Q(x) ==> (EX x. P(x)) <-> (EX x. Q(x))"
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   230
               (fn prems => [cut_facts_tac prems 1, Blast_tac 1])
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   231
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   232
val all_comm = prove_goal (the_context()) "(ALL x y. P(x,y)) <-> (ALL y x. P(x,y))"
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   233
               (fn _ => [Blast_tac 1])
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   234
val ex_comm = prove_goal (the_context()) "(EX x y. P(x,y)) <-> (EX y x. P(x,y))"
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   235
               (fn _ => [Blast_tac 1])
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   236
in
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   237
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   238
(** make simplification procedures for quantifier elimination **)
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   239
structure Quantifier1 = Quantifier1Fun(
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   240
struct
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   241
  (*abstract syntax*)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
   242
  fun dest_eq((c as Const("op =",_)) $ s $ t) = SOME(c,s,t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
   243
    | dest_eq _ = NONE;
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
   244
  fun dest_conj((c as Const("op &",_)) $ s $ t) = SOME(c,s,t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
   245
    | dest_conj _ = NONE;
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
   246
  fun dest_imp((c as Const("op -->",_)) $ s $ t) = SOME(c,s,t)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 13462
diff changeset
   247
    | dest_imp _ = NONE;
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   248
  val conj = FOLogic.conj
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   249
  val imp  = FOLogic.imp
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   250
  (*rules*)
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   251
  val iff_reflection = iff_reflection
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   252
  val iffI = iffI
12526
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   253
  val iff_trans = iff_trans
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   254
  val conjI= conjI
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   255
  val conjE= conjE
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   256
  val impI = impI
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   257
  val mp   = mp
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   258
  val uncurry = uncurry
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   259
  val exI  = exI
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   260
  val exE  = exE
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   261
  val iff_allI = iff_allI
12526
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   262
  val iff_exI = iff_exI
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   263
  val all_comm = all_comm
1b9db2581fe2 mods due to changed 1-point simprocs (quantifier1).
nipkow
parents: 12038
diff changeset
   264
  val ex_comm = ex_comm
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   265
end);
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   266
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   267
end;
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 10431
diff changeset
   268
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
   269
val defEX_regroup =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
   270
  Simplifier.simproc (Theory.sign_of (the_context ()))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
   271
    "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
   272
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   273
val defALL_regroup =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
   274
  Simplifier.simproc (Theory.sign_of (the_context ()))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13149
diff changeset
   275
    "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
   276
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   277
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   278
(*** Case splitting ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   279
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   280
bind_thm ("meta_eq_to_iff", prove_goal IFOL.thy "x==y ==> x<->y"
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   281
  (fn [prem] => [rewtac prem, rtac iffI 1, atac 1, atac 1]));
1722
bb326972ede6 Added split_inside_tac.
berghofe
parents: 1459
diff changeset
   282
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   283
structure SplitterData =
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   284
  struct
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   285
  structure Simplifier = Simplifier
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
   286
  val mk_eq          = mk_eq
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   287
  val meta_eq_to_iff = meta_eq_to_iff
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   288
  val iffD           = iffD2
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   289
  val disjE          = disjE
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   290
  val conjE          = conjE
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   291
  val exE            = exE
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   292
  val contrapos      = contrapos
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   293
  val contrapos2     = contrapos2
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   294
  val notnotD        = notnotD
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   295
  end;
1722
bb326972ede6 Added split_inside_tac.
berghofe
parents: 1459
diff changeset
   296
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   297
structure Splitter = SplitterFun(SplitterData);
1722
bb326972ede6 Added split_inside_tac.
berghofe
parents: 1459
diff changeset
   298
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   299
val split_tac        = Splitter.split_tac;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   300
val split_inside_tac = Splitter.split_inside_tac;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   301
val split_asm_tac    = Splitter.split_asm_tac;
5307
6a699d5cdef4 minor adaption for SML/NJ
oheimb
parents: 5304
diff changeset
   302
val op addsplits     = Splitter.addsplits;
6a699d5cdef4 minor adaption for SML/NJ
oheimb
parents: 5304
diff changeset
   303
val op delsplits     = Splitter.delsplits;
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   304
val Addsplits        = Splitter.Addsplits;
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   305
val Delsplits        = Splitter.Delsplits;
4325
e72cba5af6c5 addsplits now in FOL, ZF too
paulson
parents: 4203
diff changeset
   306
e72cba5af6c5 addsplits now in FOL, ZF too
paulson
parents: 4203
diff changeset
   307
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   308
(*** Standard simpsets ***)
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   309
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   310
structure Induction = InductionFun(struct val spec=IFOL.spec end);
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   311
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   312
open Induction;
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   313
5555
4b9386224084 simplified CLASIMP_DATA
oheimb
parents: 5496
diff changeset
   314
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   315
bind_thms ("meta_simps",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   316
 [triv_forall_equality,   (* prunes params *)
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   317
  True_implies_equals]);  (* prune asms `True' *)
5496
42d13691be86 Pruning of parameters and True assumptions
paulson
parents: 5307
diff changeset
   318
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   319
bind_thms ("IFOL_simps",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   320
 [refl RS P_iff_T] @ conj_simps @ disj_simps @ not_simps @
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   321
  imp_simps @ iff_simps @ quant_simps);
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   322
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   323
bind_thm ("notFalseI", int_prove_fun "~False");
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   324
bind_thms ("triv_rls", [TrueI,refl,reflexive_thm,iff_refl,notFalseI]);
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   325
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   326
fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls@prems),
9713
2c5b42311eb0 cong setup now part of Simplifier;
wenzelm
parents: 9300
diff changeset
   327
                                 atac, etac FalseE];
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   328
(*No premature instantiation of variables during simplification*)
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   329
fun   safe_solver prems = FIRST'[match_tac (triv_rls@prems),
9713
2c5b42311eb0 cong setup now part of Simplifier;
wenzelm
parents: 9300
diff changeset
   330
                                 eq_assume_tac, ematch_tac [FalseE]];
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   331
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   332
(*No simprules, but basic infastructure for simplification*)
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   333
val FOL_basic_ss = empty_ss
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   334
  setsubgoaler asm_simp_tac
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   335
  setSSolver (mk_solver "FOL safe" safe_solver)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   336
  setSolver (mk_solver "FOL unsafe" unsafe_solver)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   337
  setmksimps (mksimps mksimps_pairs)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   338
  setmkcong mk_meta_cong;
5304
c133f16febc7 the splitter is now defined as a functor
oheimb
parents: 5220
diff changeset
   339
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   340
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   341
(*intuitionistic simprules only*)
10431
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   342
val IFOL_ss = FOL_basic_ss
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   343
  addsimps (meta_simps @ IFOL_simps @ int_ex_simps @ int_all_simps)
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   344
  addsimprocs [defALL_regroup, defEX_regroup]    
bb67f704d631 FOL_basic_ss: simprocs moved to FOL_ss;
wenzelm
parents: 9889
diff changeset
   345
  addcongs [imp_cong];
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   346
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   347
bind_thms ("cla_simps",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   348
  [de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2,
12825
f1f7964ed05c new simprules and classical rules
paulson
parents: 12765
diff changeset
   349
   not_imp, not_all, not_ex, cases_simp] @
12038
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   350
  map prove_fun
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   351
   ["~(P&Q) <-> ~P | ~Q",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   352
    "P | ~P",             "~P | P",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   353
    "~ ~ P <-> P",        "(~P --> P) <-> P",
343a9888e875 proper use of bind_thm(s);
wenzelm
parents: 11771
diff changeset
   354
    "(~P <-> ~Q) <-> (P<->Q)"]);
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   355
3910
1cc9b8ab161c New simprules imp_disj1,2 and some comments
paulson
parents: 3835
diff changeset
   356
(*classical simprules too*)
4349
50403e5a44c0 Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents: 4325
diff changeset
   357
val FOL_ss = IFOL_ss addsimps (cla_simps @ cla_ex_simps @ cla_all_simps);
2074
30a65172e003 Added the de Morgan laws (incl quantifier versions) to basic simpset
paulson
parents: 2065
diff changeset
   358
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6391
diff changeset
   359
val simpsetup = [fn thy => (simpset_ref_of thy := FOL_ss; thy)];
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   360
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   361
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 5115
diff changeset
   362
(*** integration of simplifier with classical reasoner ***)
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   363
5219
924359415f09 functorized Clasimp module;
wenzelm
parents: 5115
diff changeset
   364
structure Clasimp = ClasimpFun
8472
50a653f8b8ea clasimp: include Splitter;
wenzelm
parents: 7570
diff changeset
   365
 (structure Simplifier = Simplifier and Splitter = Splitter
9851
e22db9397e17 iff declarations moved to clasimp.ML;
wenzelm
parents: 9713
diff changeset
   366
  and Classical  = Cla and Blast = Blast
11344
57b7ad51971c streamlined addIffs/delIffs, added warnings
oheimb
parents: 11232
diff changeset
   367
  val iffD1 = iffD1 val iffD2 = iffD2 val notE = notE
9851
e22db9397e17 iff declarations moved to clasimp.ML;
wenzelm
parents: 9713
diff changeset
   368
  val cla_make_elim = cla_make_elim);
4652
d24cca140eeb factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
oheimb
parents: 4649
diff changeset
   369
open Clasimp;
2633
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   370
37c0b5a7ee5d added delcongs, Delcongs, unsafe_solver, safe_solver, FOL_basic_ss,
oheimb
parents: 2601
diff changeset
   371
val FOL_css = (FOL_cs, FOL_ss);