src/HOL/simpdata.ML
author nipkow
Fri, 09 Mar 2001 19:05:48 +0100
changeset 11200 f43fa07536c0
parent 11162 9e2ec5f02217
child 11220 db536a42dfc5
permissions -rw-r--r--
arith_tac now copes with propositional reasoning as well.
clasohm@1465
     1
(*  Title:      HOL/simpdata.ML
clasohm@923
     2
    ID:         $Id$
clasohm@1465
     3
    Author:     Tobias Nipkow
clasohm@923
     4
    Copyright   1991  University of Cambridge
clasohm@923
     5
oheimb@5304
     6
Instantiation of the generic simplifier for HOL.
clasohm@923
     7
*)
clasohm@923
     8
paulson@1984
     9
section "Simplifier";
paulson@1984
    10
wenzelm@7357
    11
val [prem] = goal (the_context ()) "x==y ==> x=y";
paulson@7031
    12
by (rewtac prem);
paulson@7031
    13
by (rtac refl 1);
paulson@7031
    14
qed "meta_eq_to_obj_eq";
nipkow@4640
    15
oheimb@9023
    16
Goal "(%s. f s) = f";
oheimb@9023
    17
br refl 1;
oheimb@9023
    18
qed "eta_contract_eq";
oheimb@9023
    19
clasohm@923
    20
local
clasohm@923
    21
wenzelm@7357
    22
  fun prover s = prove_goal (the_context ()) s (fn _ => [(Blast_tac 1)]);
clasohm@923
    23
nipkow@2134
    24
in
nipkow@2134
    25
oheimb@5552
    26
(*Make meta-equalities.  The operator below is Trueprop*)
oheimb@5304
    27
nipkow@6128
    28
fun mk_meta_eq r = r RS eq_reflection;
wenzelm@9832
    29
fun safe_mk_meta_eq r = mk_meta_eq r handle Thm.THM _ => r;
oheimb@5552
    30
nipkow@6128
    31
val Eq_TrueI  = mk_meta_eq(prover  "P --> (P = True)"  RS mp);
nipkow@6128
    32
val Eq_FalseI = mk_meta_eq(prover "~P --> (P = False)" RS mp);
oheimb@5304
    33
nipkow@6128
    34
fun mk_eq th = case concl_of th of
nipkow@6128
    35
        Const("==",_)$_$_       => th
nipkow@6128
    36
    |   _$(Const("op =",_)$_$_) => mk_meta_eq th
nipkow@6128
    37
    |   _$(Const("Not",_)$_)    => th RS Eq_FalseI
nipkow@6128
    38
    |   _                       => th RS Eq_TrueI;
nipkow@6128
    39
(* last 2 lines requires all formulae to be of the from Trueprop(.) *)
oheimb@5552
    40
nipkow@6128
    41
fun mk_eq_True r = Some(r RS meta_eq_to_obj_eq RS Eq_TrueI);
nipkow@6128
    42
wenzelm@9713
    43
(*Congruence rules for = (instead of ==)*)
nipkow@6128
    44
fun mk_meta_cong rl =
nipkow@6128
    45
  standard(mk_meta_eq(replicate (nprems_of rl) meta_eq_to_obj_eq MRS rl))
nipkow@6128
    46
  handle THM _ =>
nipkow@6128
    47
  error("Premises and conclusion of congruence rules must be =-equalities");
nipkow@3896
    48
nipkow@5975
    49
val not_not = prover "(~ ~ P) = P";
clasohm@923
    50
nipkow@5975
    51
val simp_thms = [not_not] @ map prover
paulson@2082
    52
 [ "(x=x) = True",
nipkow@5975
    53
   "(~True) = False", "(~False) = True",
paulson@2082
    54
   "(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))",
nipkow@4640
    55
   "(True=P) = P", "(P=True) = P", "(False=P) = (~P)", "(P=False) = (~P)",
wenzelm@9713
    56
   "(True --> P) = P", "(False --> P) = True",
paulson@2082
    57
   "(P --> True) = True", "(P --> P) = True",
paulson@2082
    58
   "(P --> False) = (~P)", "(P --> ~P) = (~P)",
wenzelm@9713
    59
   "(P & True) = P", "(True & P) = P",
nipkow@2800
    60
   "(P & False) = False", "(False & P) = False",
nipkow@2800
    61
   "(P & P) = P", "(P & (P & Q)) = (P & Q)",
paulson@3913
    62
   "(P & ~P) = False",    "(~P & P) = False",
wenzelm@9713
    63
   "(P | True) = True", "(True | P) = True",
nipkow@2800
    64
   "(P | False) = P", "(False | P) = P",
nipkow@2800
    65
   "(P | P) = P", "(P | (P | Q)) = (P | Q)",
paulson@3913
    66
   "(P | ~P) = True",    "(~P | P) = True",
paulson@2082
    67
   "((~P) = (~Q)) = (P=Q)",
wenzelm@9713
    68
   "(!x. P) = P", "(? x. P) = P", "? x. x=t", "? x. t=x",
paulson@4351
    69
(*two needed for the one-point-rule quantifier simplification procs*)
wenzelm@9713
    70
   "(? x. x=t & P(x)) = P(t)",          (*essential for termination!!*)
paulson@4351
    71
   "(! x. t=x --> P(x)) = P(t)" ];      (*covers a stray case*)
clasohm@923
    72
nipkow@9875
    73
val imp_cong = standard(impI RSN
wenzelm@7357
    74
    (2, prove_goal (the_context ()) "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))"
nipkow@9875
    75
        (fn _=> [(Blast_tac 1)]) RS mp RS mp));
paulson@1922
    76
paulson@1948
    77
(*Miniscoping: pushing in existential quantifiers*)
wenzelm@7648
    78
val ex_simps = map prover
wenzelm@3842
    79
                ["(EX x. P x & Q)   = ((EX x. P x) & Q)",
wenzelm@3842
    80
                 "(EX x. P & Q x)   = (P & (EX x. Q x))",
wenzelm@3842
    81
                 "(EX x. P x | Q)   = ((EX x. P x) | Q)",
wenzelm@3842
    82
                 "(EX x. P | Q x)   = (P | (EX x. Q x))",
wenzelm@3842
    83
                 "(EX x. P x --> Q) = ((ALL x. P x) --> Q)",
wenzelm@3842
    84
                 "(EX x. P --> Q x) = (P --> (EX x. Q x))"];
paulson@1948
    85
paulson@1948
    86
(*Miniscoping: pushing in universal quantifiers*)
paulson@1948
    87
val all_simps = map prover
wenzelm@3842
    88
                ["(ALL x. P x & Q)   = ((ALL x. P x) & Q)",
wenzelm@3842
    89
                 "(ALL x. P & Q x)   = (P & (ALL x. Q x))",
wenzelm@3842
    90
                 "(ALL x. P x | Q)   = ((ALL x. P x) | Q)",
wenzelm@3842
    91
                 "(ALL x. P | Q x)   = (P | (ALL x. Q x))",
wenzelm@3842
    92
                 "(ALL x. P x --> Q) = ((EX x. P x) --> Q)",
wenzelm@3842
    93
                 "(ALL x. P --> Q x) = (P --> (ALL x. Q x))"];
paulson@1948
    94
clasohm@923
    95
paulson@2022
    96
(* elimination of existential quantifiers in assumptions *)
clasohm@923
    97
clasohm@923
    98
val ex_all_equiv =
wenzelm@7357
    99
  let val lemma1 = prove_goal (the_context ())
clasohm@923
   100
        "(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)"
clasohm@923
   101
        (fn prems => [resolve_tac prems 1, etac exI 1]);
wenzelm@7357
   102
      val lemma2 = prove_goalw (the_context ()) [Ex_def]
clasohm@923
   103
        "(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)"
paulson@7031
   104
        (fn prems => [(REPEAT(resolve_tac prems 1))])
clasohm@923
   105
  in equal_intr lemma1 lemma2 end;
clasohm@923
   106
clasohm@923
   107
end;
clasohm@923
   108
wenzelm@7648
   109
bind_thms ("ex_simps", ex_simps);
wenzelm@7648
   110
bind_thms ("all_simps", all_simps);
berghofe@7711
   111
bind_thm ("not_not", not_not);
nipkow@9875
   112
bind_thm ("imp_cong", imp_cong);
wenzelm@7648
   113
nipkow@3654
   114
(* Elimination of True from asumptions: *)
nipkow@3654
   115
wenzelm@7357
   116
val True_implies_equals = prove_goal (the_context ())
nipkow@3654
   117
 "(True ==> PROP P) == PROP P"
paulson@7031
   118
(fn _ => [rtac equal_intr_rule 1, atac 2,
nipkow@3654
   119
          METAHYPS (fn prems => resolve_tac prems 1) 1,
nipkow@3654
   120
          rtac TrueI 1]);
nipkow@3654
   121
wenzelm@7357
   122
fun prove nm thm  = qed_goal nm (the_context ()) thm (fn _ => [(Blast_tac 1)]);
clasohm@923
   123
paulson@9511
   124
prove "eq_commute" "(a=b) = (b=a)";
paulson@7623
   125
prove "eq_left_commute" "(P=(Q=R)) = (Q=(P=R))";
paulson@7623
   126
prove "eq_assoc" "((P=Q)=R) = (P=(Q=R))";
paulson@7623
   127
val eq_ac = [eq_commute, eq_left_commute, eq_assoc];
paulson@7623
   128
paulson@9511
   129
prove "neq_commute" "(a~=b) = (b~=a)";
paulson@9511
   130
clasohm@923
   131
prove "conj_commute" "(P&Q) = (Q&P)";
clasohm@923
   132
prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))";
clasohm@923
   133
val conj_comms = [conj_commute, conj_left_commute];
nipkow@2134
   134
prove "conj_assoc" "((P&Q)&R) = (P&(Q&R))";
clasohm@923
   135
paulson@1922
   136
prove "disj_commute" "(P|Q) = (Q|P)";
paulson@1922
   137
prove "disj_left_commute" "(P|(Q|R)) = (Q|(P|R))";
paulson@1922
   138
val disj_comms = [disj_commute, disj_left_commute];
nipkow@2134
   139
prove "disj_assoc" "((P|Q)|R) = (P|(Q|R))";
paulson@1922
   140
clasohm@923
   141
prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)";
clasohm@923
   142
prove "conj_disj_distribR" "((P|Q)&R) = (P&R | Q&R)";
nipkow@1485
   143
paulson@1892
   144
prove "disj_conj_distribL" "(P|(Q&R)) = ((P|Q) & (P|R))";
paulson@1892
   145
prove "disj_conj_distribR" "((P&Q)|R) = ((P|R) & (Q|R))";
paulson@1892
   146
nipkow@2134
   147
prove "imp_conjR" "(P --> (Q&R)) = ((P-->Q) & (P-->R))";
nipkow@2134
   148
prove "imp_conjL" "((P&Q) -->R)  = (P --> (Q --> R))";
nipkow@2134
   149
prove "imp_disjL" "((P|Q) --> R) = ((P-->R)&(Q-->R))";
paulson@1892
   150
paulson@3448
   151
(*These two are specialized, but imp_disj_not1 is useful in Auth/Yahalom.ML*)
paulson@8114
   152
prove "imp_disj_not1" "(P --> Q | R) = (~Q --> P --> R)";
paulson@8114
   153
prove "imp_disj_not2" "(P --> Q | R) = (~R --> P --> Q)";
paulson@3448
   154
paulson@3904
   155
prove "imp_disj1" "((P-->Q)|R) = (P--> Q|R)";
paulson@3904
   156
prove "imp_disj2" "(Q|(P-->R)) = (P--> Q|R)";
paulson@3904
   157
nipkow@1485
   158
prove "de_Morgan_disj" "(~(P | Q)) = (~P & ~Q)";
nipkow@1485
   159
prove "de_Morgan_conj" "(~(P & Q)) = (~P | ~Q)";
paulson@3446
   160
prove "not_imp" "(~(P --> Q)) = (P & ~Q)";
paulson@1922
   161
prove "not_iff" "(P~=Q) = (P = (~Q))";
oheimb@4743
   162
prove "disj_not1" "(~P | Q) = (P --> Q)";
oheimb@4743
   163
prove "disj_not2" "(P | ~Q) = (Q --> P)"; (* changes orientation :-( *)
nipkow@5975
   164
prove "imp_conv_disj" "(P --> Q) = ((~P) | Q)";
nipkow@5975
   165
nipkow@5975
   166
prove "iff_conv_conj_imp" "(P = Q) = ((P --> Q) & (Q --> P))";
nipkow@5975
   167
nipkow@1485
   168
wenzelm@9713
   169
(*Avoids duplication of subgoals after split_if, when the true and false
wenzelm@9713
   170
  cases boil down to the same thing.*)
nipkow@2134
   171
prove "cases_simp" "((P --> Q) & (~P --> Q)) = Q";
nipkow@2134
   172
wenzelm@3842
   173
prove "not_all" "(~ (! x. P(x))) = (? x.~P(x))";
paulson@1922
   174
prove "imp_all" "((! x. P x) --> Q) = (? x. P x --> Q)";
wenzelm@3842
   175
prove "not_ex"  "(~ (? x. P(x))) = (! x.~P(x))";
paulson@1922
   176
prove "imp_ex" "((? x. P x) --> Q) = (! x. P x --> Q)";
oheimb@1660
   177
nipkow@1655
   178
prove "ex_disj_distrib" "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))";
nipkow@1655
   179
prove "all_conj_distrib" "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))";
nipkow@1655
   180
nipkow@2134
   181
(* '&' congruence rule: not included by default!
nipkow@2134
   182
   May slow rewrite proofs down by as much as 50% *)
nipkow@2134
   183
wenzelm@9713
   184
let val th = prove_goal (the_context ())
nipkow@2134
   185
                "(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))"
paulson@7031
   186
                (fn _=> [(Blast_tac 1)])
nipkow@2134
   187
in  bind_thm("conj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
nipkow@2134
   188
wenzelm@9713
   189
let val th = prove_goal (the_context ())
nipkow@2134
   190
                "(Q=Q')--> (Q'--> (P=P'))--> ((P&Q) = (P'&Q'))"
paulson@7031
   191
                (fn _=> [(Blast_tac 1)])
nipkow@2134
   192
in  bind_thm("rev_conj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
nipkow@2134
   193
nipkow@2134
   194
(* '|' congruence rule: not included by default! *)
nipkow@2134
   195
wenzelm@9713
   196
let val th = prove_goal (the_context ())
nipkow@2134
   197
                "(P=P')--> (~P'--> (Q=Q'))--> ((P|Q) = (P'|Q'))"
paulson@7031
   198
                (fn _=> [(Blast_tac 1)])
nipkow@2134
   199
in  bind_thm("disj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
nipkow@2134
   200
nipkow@2134
   201
prove "eq_sym_conv" "(x=y) = (y=x)";
nipkow@2134
   202
paulson@5278
   203
paulson@5278
   204
(** if-then-else rules **)
paulson@5278
   205
paulson@7031
   206
Goalw [if_def] "(if True then x else y) = x";
paulson@7031
   207
by (Blast_tac 1);
paulson@7031
   208
qed "if_True";
nipkow@2134
   209
paulson@7031
   210
Goalw [if_def] "(if False then x else y) = y";
paulson@7031
   211
by (Blast_tac 1);
paulson@7031
   212
qed "if_False";
nipkow@2134
   213
paulson@7127
   214
Goalw [if_def] "P ==> (if P then x else y) = x";
paulson@7031
   215
by (Blast_tac 1);
paulson@7031
   216
qed "if_P";
oheimb@5304
   217
paulson@7127
   218
Goalw [if_def] "~P ==> (if P then x else y) = y";
paulson@7031
   219
by (Blast_tac 1);
paulson@7031
   220
qed "if_not_P";
nipkow@2134
   221
paulson@7031
   222
Goal "P(if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))";
paulson@7031
   223
by (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1);
paulson@7031
   224
by (stac if_P 2);
paulson@7031
   225
by (stac if_not_P 1);
paulson@7031
   226
by (ALLGOALS (Blast_tac));
paulson@7031
   227
qed "split_if";
paulson@7031
   228
paulson@7031
   229
Goal "P(if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))";
paulson@7031
   230
by (stac split_if 1);
paulson@7031
   231
by (Blast_tac 1);
paulson@7031
   232
qed "split_if_asm";
nipkow@2134
   233
wenzelm@9384
   234
bind_thms ("if_splits", [split_if, split_if_asm]);
wenzelm@9384
   235
oheimb@11003
   236
bind_thm ("if_def2", read_instantiate [("P","\\<lambda>x. x")] split_if);
oheimb@11003
   237
paulson@7031
   238
Goal "(if c then x else x) = x";
paulson@7031
   239
by (stac split_if 1);
paulson@7031
   240
by (Blast_tac 1);
paulson@7031
   241
qed "if_cancel";
oheimb@5304
   242
paulson@7031
   243
Goal "(if x = y then y else x) = x";
paulson@7031
   244
by (stac split_if 1);
paulson@7031
   245
by (Blast_tac 1);
paulson@7031
   246
qed "if_eq_cancel";
oheimb@5304
   247
paulson@4769
   248
(*This form is useful for expanding IFs on the RIGHT of the ==> symbol*)
paulson@7127
   249
Goal "(if P then Q else R) = ((P-->Q) & (~P-->R))";
paulson@7031
   250
by (rtac split_if 1);
paulson@7031
   251
qed "if_bool_eq_conj";
paulson@4769
   252
paulson@4769
   253
(*And this form is useful for expanding IFs on the LEFT*)
paulson@7031
   254
Goal "(if P then Q else R) = ((P&Q) | (~P&R))";
paulson@7031
   255
by (stac split_if 1);
paulson@7031
   256
by (Blast_tac 1);
paulson@7031
   257
qed "if_bool_eq_disj";
nipkow@2134
   258
paulson@4351
   259
paulson@4351
   260
(*** make simplification procedures for quantifier elimination ***)
paulson@4351
   261
wenzelm@9851
   262
structure Quantifier1 = Quantifier1Fun
wenzelm@9851
   263
(struct
paulson@4351
   264
  (*abstract syntax*)
paulson@4351
   265
  fun dest_eq((c as Const("op =",_)) $ s $ t) = Some(c,s,t)
paulson@4351
   266
    | dest_eq _ = None;
paulson@4351
   267
  fun dest_conj((c as Const("op &",_)) $ s $ t) = Some(c,s,t)
paulson@4351
   268
    | dest_conj _ = None;
paulson@4351
   269
  val conj = HOLogic.conj
paulson@4351
   270
  val imp  = HOLogic.imp
paulson@4351
   271
  (*rules*)
paulson@4351
   272
  val iff_reflection = eq_reflection
paulson@4351
   273
  val iffI = iffI
paulson@4351
   274
  val sym  = sym
paulson@4351
   275
  val conjI= conjI
paulson@4351
   276
  val conjE= conjE
paulson@4351
   277
  val impI = impI
paulson@4351
   278
  val impE = impE
paulson@4351
   279
  val mp   = mp
paulson@4351
   280
  val exI  = exI
paulson@4351
   281
  val exE  = exE
paulson@4351
   282
  val allI = allI
paulson@4351
   283
  val allE = allE
paulson@4351
   284
end);
paulson@4351
   285
nipkow@4320
   286
local
nipkow@4320
   287
val ex_pattern =
wenzelm@7357
   288
  Thm.read_cterm (Theory.sign_of (the_context ())) ("EX x. P(x) & Q(x)",HOLogic.boolT)
paulson@3913
   289
nipkow@4320
   290
val all_pattern =
wenzelm@7357
   291
  Thm.read_cterm (Theory.sign_of (the_context ())) ("ALL x. P(x) & P'(x) --> Q(x)",HOLogic.boolT)
nipkow@4320
   292
nipkow@4320
   293
in
nipkow@4320
   294
val defEX_regroup =
nipkow@4320
   295
  mk_simproc "defined EX" [ex_pattern] Quantifier1.rearrange_ex;
nipkow@4320
   296
val defALL_regroup =
nipkow@4320
   297
  mk_simproc "defined ALL" [all_pattern] Quantifier1.rearrange_all;
nipkow@4320
   298
end;
paulson@3913
   299
paulson@4351
   300
paulson@4351
   301
(*** Case splitting ***)
paulson@3913
   302
oheimb@5304
   303
structure SplitterData =
oheimb@5304
   304
  struct
oheimb@5304
   305
  structure Simplifier = Simplifier
oheimb@5552
   306
  val mk_eq          = mk_eq
oheimb@5304
   307
  val meta_eq_to_iff = meta_eq_to_obj_eq
oheimb@5304
   308
  val iffD           = iffD2
oheimb@5304
   309
  val disjE          = disjE
oheimb@5304
   310
  val conjE          = conjE
oheimb@5304
   311
  val exE            = exE
paulson@10231
   312
  val contrapos      = contrapos_nn
paulson@10231
   313
  val contrapos2     = contrapos_pp
oheimb@5304
   314
  val notnotD        = notnotD
oheimb@5304
   315
  end;
oheimb@2263
   316
oheimb@5304
   317
structure Splitter = SplitterFun(SplitterData);
oheimb@2263
   318
oheimb@5304
   319
val split_tac        = Splitter.split_tac;
oheimb@5304
   320
val split_inside_tac = Splitter.split_inside_tac;
oheimb@5304
   321
val split_asm_tac    = Splitter.split_asm_tac;
oheimb@5307
   322
val op addsplits     = Splitter.addsplits;
oheimb@5307
   323
val op delsplits     = Splitter.delsplits;
oheimb@5304
   324
val Addsplits        = Splitter.Addsplits;
oheimb@5304
   325
val Delsplits        = Splitter.Delsplits;
oheimb@4718
   326
nipkow@2134
   327
(*In general it seems wrong to add distributive laws by default: they
nipkow@2134
   328
  might cause exponential blow-up.  But imp_disjL has been in for a while
wenzelm@9713
   329
  and cannot be removed without affecting existing proofs.  Moreover,
nipkow@2134
   330
  rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the
nipkow@2134
   331
  grounds that it allows simplification of R in the two cases.*)
nipkow@2134
   332
nipkow@2134
   333
val mksimps_pairs =
nipkow@2134
   334
  [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
nipkow@2134
   335
   ("All", [spec]), ("True", []), ("False", []),
paulson@4769
   336
   ("If", [if_bool_eq_conj RS iffD1])];
nipkow@1758
   337
oheimb@5552
   338
(* ###FIXME: move to Provers/simplifier.ML
oheimb@5304
   339
val mk_atomize:      (string * thm list) list -> thm -> thm list
oheimb@5304
   340
*)
oheimb@5552
   341
(* ###FIXME: move to Provers/simplifier.ML *)
oheimb@5304
   342
fun mk_atomize pairs =
oheimb@5304
   343
  let fun atoms th =
oheimb@5304
   344
        (case concl_of th of
oheimb@5304
   345
           Const("Trueprop",_) $ p =>
oheimb@5304
   346
             (case head_of p of
oheimb@5304
   347
                Const(a,_) =>
oheimb@5304
   348
                  (case assoc(pairs,a) of
oheimb@5304
   349
                     Some(rls) => flat (map atoms ([th] RL rls))
oheimb@5304
   350
                   | None => [th])
oheimb@5304
   351
              | _ => [th])
oheimb@5304
   352
         | _ => [th])
oheimb@5304
   353
  in atoms end;
oheimb@5304
   354
oheimb@11162
   355
fun mksimps pairs = (map mk_eq o mk_atomize pairs o forall_elim_vars_safe);
oheimb@5304
   356
nipkow@7570
   357
fun unsafe_solver_tac prems =
nipkow@7570
   358
  FIRST'[resolve_tac(reflexive_thm::TrueI::refl::prems), atac, etac FalseE];
nipkow@7570
   359
val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac;
nipkow@7570
   360
oheimb@2636
   361
(*No premature instantiation of variables during simplification*)
nipkow@7570
   362
fun safe_solver_tac prems =
nipkow@7570
   363
  FIRST'[match_tac(reflexive_thm::TrueI::refl::prems),
nipkow@7570
   364
         eq_assume_tac, ematch_tac [FalseE]];
nipkow@7570
   365
val safe_solver = mk_solver "HOL safe" safe_solver_tac;
oheimb@2443
   366
wenzelm@9713
   367
val HOL_basic_ss =
wenzelm@9713
   368
  empty_ss setsubgoaler asm_simp_tac
wenzelm@9713
   369
    setSSolver safe_solver
wenzelm@9713
   370
    setSolver unsafe_solver
wenzelm@9713
   371
    setmksimps (mksimps mksimps_pairs)
wenzelm@9713
   372
    setmkeqTrue mk_eq_True
wenzelm@9713
   373
    setmkcong mk_meta_cong;
oheimb@2443
   374
wenzelm@9713
   375
val HOL_ss =
wenzelm@9713
   376
    HOL_basic_ss addsimps
paulson@3446
   377
     ([triv_forall_equality, (* prunes params *)
nipkow@3654
   378
       True_implies_equals, (* prune asms `True' *)
oheimb@9023
   379
       eta_contract_eq, (* prunes eta-expansions *)
oheimb@4718
   380
       if_True, if_False, if_cancel, if_eq_cancel,
oheimb@5304
   381
       imp_disjL, conj_assoc, disj_assoc,
paulson@3904
   382
       de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, not_imp,
paulson@9969
   383
       disj_not1, not_all, not_ex, cases_simp, some_eq_trivial, some_sym_eq_trivial,
paulson@8955
   384
       thm"plus_ac0.zero", thm"plus_ac0_zero_right"]
paulson@3446
   385
     @ ex_simps @ all_simps @ simp_thms)
nipkow@4032
   386
     addsimprocs [defALL_regroup,defEX_regroup]
wenzelm@4744
   387
     addcongs [imp_cong]
nipkow@4830
   388
     addsplits [split_if];
paulson@2082
   389
wenzelm@11034
   390
fun hol_simplify rews = Simplifier.full_simplify (HOL_basic_ss addsimps rews);
wenzelm@11034
   391
fun hol_rewrite_cterm rews =
wenzelm@11034
   392
  #2 o Thm.dest_comb o #prop o Thm.crep_thm o Simplifier.full_rewrite (HOL_basic_ss addsimps rews);
wenzelm@11034
   393
wenzelm@11034
   394
paulson@6293
   395
(*Simplifies x assuming c and y assuming ~c*)
paulson@6293
   396
val prems = Goalw [if_def]
paulson@6293
   397
  "[| b=c; c ==> x=u; ~c ==> y=v |] ==> \
paulson@6293
   398
\  (if b then x else y) = (if c then u else v)";
paulson@6293
   399
by (asm_simp_tac (HOL_ss addsimps prems) 1);
paulson@6293
   400
qed "if_cong";
paulson@6293
   401
paulson@7127
   402
(*Prevents simplification of x and y: faster and allows the execution
paulson@7127
   403
  of functional programs. NOW THE DEFAULT.*)
paulson@7031
   404
Goal "b=c ==> (if b then x else y) = (if c then x else y)";
paulson@7031
   405
by (etac arg_cong 1);
paulson@7031
   406
qed "if_weak_cong";
paulson@6293
   407
paulson@6293
   408
(*Prevents simplification of t: much faster*)
paulson@7031
   409
Goal "a = b ==> (let x=a in t(x)) = (let x=b in t(x))";
paulson@7031
   410
by (etac arg_cong 1);
paulson@7031
   411
qed "let_weak_cong";
paulson@6293
   412
paulson@7031
   413
Goal "f(if c then x else y) = (if c then f x else f y)";
paulson@7031
   414
by (simp_tac (HOL_ss setloop (split_tac [split_if])) 1);
paulson@7031
   415
qed "if_distrib";
nipkow@1655
   416
paulson@4327
   417
(*For expand_case_tac*)
paulson@7584
   418
val prems = Goal "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)";
paulson@2948
   419
by (case_tac "P" 1);
paulson@2948
   420
by (ALLGOALS (asm_simp_tac (HOL_ss addsimps prems)));
paulson@7584
   421
qed "expand_case";
paulson@2948
   422
paulson@4327
   423
(*Used in Auth proofs.  Typically P contains Vars that become instantiated
paulson@4327
   424
  during unification.*)
paulson@2948
   425
fun expand_case_tac P i =
paulson@2948
   426
    res_inst_tac [("P",P)] expand_case i THEN
wenzelm@9713
   427
    Simp_tac (i+1) THEN
paulson@2948
   428
    Simp_tac i;
paulson@2948
   429
paulson@7584
   430
(*This lemma restricts the effect of the rewrite rule u=v to the left-hand
paulson@7584
   431
  side of an equality.  Used in {Integ,Real}/simproc.ML*)
paulson@7584
   432
Goal "x=y ==> (x=z) = (y=z)";
paulson@7584
   433
by (asm_simp_tac HOL_ss 1);
paulson@7584
   434
qed "restrict_to_left";
paulson@2948
   435
wenzelm@7357
   436
(* default simpset *)
wenzelm@9713
   437
val simpsetup =
wenzelm@9713
   438
  [fn thy => (simpset_ref_of thy := HOL_ss addcongs [if_weak_cong]; thy)];
berghofe@3615
   439
oheimb@4652
   440
wenzelm@5219
   441
(*** integration of simplifier with classical reasoner ***)
oheimb@2636
   442
wenzelm@5219
   443
structure Clasimp = ClasimpFun
wenzelm@8473
   444
 (structure Simplifier = Simplifier and Splitter = Splitter
wenzelm@9851
   445
  and Classical  = Classical and Blast = Blast
wenzelm@9851
   446
  val dest_Trueprop = HOLogic.dest_Trueprop
wenzelm@9851
   447
  val iff_const = HOLogic.eq_const HOLogic.boolT
wenzelm@9851
   448
  val not_const = HOLogic.not_const
wenzelm@9851
   449
  val notE = notE val iffD1 = iffD1 val iffD2 = iffD2
wenzelm@9851
   450
  val cla_make_elim = cla_make_elim);
oheimb@4652
   451
open Clasimp;
oheimb@2636
   452
oheimb@2636
   453
val HOL_css = (HOL_cs, HOL_ss);
nipkow@5975
   454
nipkow@5975
   455
wenzelm@8641
   456
nipkow@5975
   457
(*** A general refutation procedure ***)
wenzelm@9713
   458
nipkow@5975
   459
(* Parameters:
nipkow@5975
   460
nipkow@5975
   461
   test: term -> bool
nipkow@5975
   462
   tests if a term is at all relevant to the refutation proof;
nipkow@5975
   463
   if not, then it can be discarded. Can improve performance,
nipkow@5975
   464
   esp. if disjunctions can be discarded (no case distinction needed!).
nipkow@5975
   465
nipkow@5975
   466
   prep_tac: int -> tactic
nipkow@5975
   467
   A preparation tactic to be applied to the goal once all relevant premises
nipkow@5975
   468
   have been moved to the conclusion.
nipkow@5975
   469
nipkow@5975
   470
   ref_tac: int -> tactic
nipkow@5975
   471
   the actual refutation tactic. Should be able to deal with goals
nipkow@5975
   472
   [| A1; ...; An |] ==> False
wenzelm@9876
   473
   where the Ai are atomic, i.e. no top-level &, | or EX
nipkow@5975
   474
*)
nipkow@5975
   475
nipkow@5975
   476
fun refute_tac test prep_tac ref_tac =
nipkow@5975
   477
  let val nnf_simps =
nipkow@5975
   478
        [imp_conv_disj,iff_conv_conj_imp,de_Morgan_disj,de_Morgan_conj,
nipkow@5975
   479
         not_all,not_ex,not_not];
nipkow@5975
   480
      val nnf_simpset =
nipkow@5975
   481
        empty_ss setmkeqTrue mk_eq_True
nipkow@5975
   482
                 setmksimps (mksimps mksimps_pairs)
nipkow@5975
   483
                 addsimps nnf_simps;
nipkow@5975
   484
      val prem_nnf_tac = full_simp_tac nnf_simpset;
nipkow@5975
   485
nipkow@5975
   486
      val refute_prems_tac =
nipkow@5975
   487
        REPEAT(eresolve_tac [conjE, exE] 1 ORELSE
nipkow@5975
   488
               filter_prems_tac test 1 ORELSE
paulson@6301
   489
               etac disjE 1) THEN
nipkow@11200
   490
        ((etac notE 1 THEN eq_assume_tac 1) ORELSE
nipkow@11200
   491
         ref_tac 1);
nipkow@5975
   492
  in EVERY'[TRY o filter_prems_tac test,
nipkow@6128
   493
            DETERM o REPEAT o etac rev_mp, prep_tac, rtac ccontr, prem_nnf_tac,
nipkow@5975
   494
            SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)]
nipkow@5975
   495
  end;