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