eliminated out-of-scope proofs (cf. theory IFOL and FOL);
authorwenzelm
Sat, 15 Mar 2008 22:07:28 +0100
changeset 26288 89b9f7c18631
parent 26287 df8e5362cff9
child 26289 9d2c375e242b
eliminated out-of-scope proofs (cf. theory IFOL and FOL); proper antiquotations;
src/FOL/simpdata.ML
--- a/src/FOL/simpdata.ML	Sat Mar 15 22:07:26 2008 +0100
+++ b/src/FOL/simpdata.ML	Sat Mar 15 22:07:28 2008 +0100
@@ -6,70 +6,11 @@
 Simplification data for FOL.
 *)
 
-(*** Rewrite rules ***)
-
-fun int_prove_fun s =
- (writeln s;
-  prove_goal (theory "IFOL") s
-   (fn prems => [ (cut_facts_tac prems 1),
-                  (IntPr.fast_tac 1) ]));
-
-bind_thms ("conj_simps", map int_prove_fun
- ["P & True <-> P",      "True & P <-> P",
-  "P & False <-> False", "False & P <-> False",
-  "P & P <-> P", "P & P & Q <-> P & Q",
-  "P & ~P <-> False",    "~P & P <-> False",
-  "(P & Q) & R <-> P & (Q & R)"]);
-
-bind_thms ("disj_simps", map int_prove_fun
- ["P | True <-> True",  "True | P <-> True",
-  "P | False <-> P",    "False | P <-> P",
-  "P | P <-> P", "P | P | Q <-> P | Q",
-  "(P | Q) | R <-> P | (Q | R)"]);
-
-bind_thms ("not_simps", map int_prove_fun
- ["~(P|Q)  <-> ~P & ~Q",
-  "~ False <-> True",   "~ True <-> False"]);
-
-bind_thms ("imp_simps", map int_prove_fun
- ["(P --> False) <-> ~P",       "(P --> True) <-> True",
-  "(False --> P) <-> True",     "(True --> P) <-> P",
-  "(P --> P) <-> True",         "(P --> ~P) <-> ~P"]);
-
-bind_thms ("iff_simps", map int_prove_fun
- ["(True <-> P) <-> P",         "(P <-> True) <-> P",
-  "(P <-> P) <-> True",
-  "(False <-> P) <-> ~P",       "(P <-> False) <-> ~P"]);
-
-(*The x=t versions are needed for the simplification procedures*)
-bind_thms ("quant_simps", map int_prove_fun
- ["(ALL x. P) <-> P",
-  "(ALL x. x=t --> P(x)) <-> P(t)",
-  "(ALL x. t=x --> P(x)) <-> P(t)",
-  "(EX x. P) <-> P",
-  "EX x. x=t", "EX x. t=x",
-  "(EX x. x=t & P(x)) <-> P(t)",
-  "(EX x. t=x & P(x)) <-> P(t)"]);
-
-(*These are NOT supplied by default!*)
-bind_thms ("distrib_simps", map int_prove_fun
- ["P & (Q | R) <-> P&Q | P&R",
-  "(Q | R) & P <-> Q&P | R&P",
-  "(P | Q --> R) <-> (P --> R) & (Q --> R)"]);
-
-(** Conversion into rewrite rules **)
-
-bind_thm ("P_iff_F", int_prove_fun "~P ==> (P <-> False)");
-bind_thm ("iff_reflection_F", P_iff_F RS iff_reflection);
-
-bind_thm ("P_iff_T", int_prove_fun "P ==> (P <-> True)");
-bind_thm ("iff_reflection_T", P_iff_T RS iff_reflection);
-
 (*Make meta-equalities.  The operator below is Trueprop*)
 
 fun mk_meta_eq th = case concl_of th of
-    _ $ (Const("op =",_)$_$_)   => th RS eq_reflection
-  | _ $ (Const("op <->",_)$_$_) => th RS iff_reflection
+    _ $ (Const("op =",_)$_$_)   => th RS @{thm eq_reflection}
+  | _ $ (Const("op <->",_)$_$_) => th RS @{thm iff_reflection}
   | _                           =>
   error("conclusion must be a =-equality or <->");;
 
@@ -77,13 +18,13 @@
     Const("==",_)$_$_           => th
   | _ $ (Const("op =",_)$_$_)   => mk_meta_eq th
   | _ $ (Const("op <->",_)$_$_) => mk_meta_eq th
-  | _ $ (Const("Not",_)$_)      => th RS iff_reflection_F
-  | _                           => th RS iff_reflection_T;
+  | _ $ (Const("Not",_)$_)      => th RS @{thm iff_reflection_F}
+  | _                           => th RS @{thm iff_reflection_T};
 
 (*Replace premises x=y, X<->Y by X==Y*)
 val mk_meta_prems =
     rule_by_tactic
-      (REPEAT_FIRST (resolve_tac [meta_eq_to_obj_eq, thm "def_imp_iff"]));
+      (REPEAT_FIRST (resolve_tac [@{thm meta_eq_to_obj_eq}, @{thm def_imp_iff}]));
 
 (*Congruence rules for = or <-> (instead of ==)*)
 fun mk_meta_cong rl =
@@ -92,8 +33,8 @@
   error("Premises and conclusion of congruence rules must use =-equality or <->");
 
 val mksimps_pairs =
-  [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
-   ("All", [spec]), ("True", []), ("False", [])];
+  [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]),
+   ("All", [@{thm spec}]), ("True", []), ("False", [])];
 
 (* ###FIXME: move to simplifier.ML
 val mk_atomize:      (string * thm list) list -> thm -> thm list
@@ -114,99 +55,6 @@
 
 fun mksimps pairs = (map mk_eq o mk_atomize pairs o gen_all);
 
-(*** Classical laws ***)
-
-fun prove_fun s =
- (writeln s;
-  prove_goal (the_context ()) s
-   (fn prems => [ (cut_facts_tac prems 1),
-                  (Cla.fast_tac FOL_cs 1) ]));
-
-(*Avoids duplication of subgoals after expand_if, when the true and false
-  cases boil down to the same thing.*)
-bind_thm ("cases_simp", prove_fun "(P --> Q) & (~P --> Q) <-> Q");
-
-
-(*** Miniscoping: pushing quantifiers in
-     We do NOT distribute of ALL over &, or dually that of EX over |
-     Baaz and Leitsch, On Skolemization and Proof Complexity (1994)
-     show that this step can increase proof length!
-***)
-
-(*existential miniscoping*)
-bind_thms ("int_ex_simps", map int_prove_fun
- ["(EX x. P(x) & Q) <-> (EX x. P(x)) & Q",
-  "(EX x. P & Q(x)) <-> P & (EX x. Q(x))",
-  "(EX x. P(x) | Q) <-> (EX x. P(x)) | Q",
-  "(EX x. P | Q(x)) <-> P | (EX x. Q(x))"]);
-
-(*classical rules*)
-bind_thms ("cla_ex_simps", map prove_fun
- ["(EX x. P(x) --> Q) <-> (ALL x. P(x)) --> Q",
-  "(EX x. P --> Q(x)) <-> P --> (EX x. Q(x))"]);
-
-bind_thms ("ex_simps", int_ex_simps @ cla_ex_simps);
-
-(*universal miniscoping*)
-bind_thms ("int_all_simps", map int_prove_fun
- ["(ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q",
-  "(ALL x. P & Q(x)) <-> P & (ALL x. Q(x))",
-  "(ALL x. P(x) --> Q) <-> (EX x. P(x)) --> Q",
-  "(ALL x. P --> Q(x)) <-> P --> (ALL x. Q(x))"]);
-
-(*classical rules*)
-bind_thms ("cla_all_simps", map prove_fun
- ["(ALL x. P(x) | Q) <-> (ALL x. P(x)) | Q",
-  "(ALL x. P | Q(x)) <-> P | (ALL x. Q(x))"]);
-
-bind_thms ("all_simps", int_all_simps @ cla_all_simps);
-
-
-(*** Named rewrite rules proved for IFOL ***)
-
-fun int_prove nm thm  = qed_goal nm (theory "IFOL") thm
-    (fn prems => [ (cut_facts_tac prems 1),
-                   (IntPr.fast_tac 1) ]);
-
-fun prove nm thm  = qed_goal nm (the_context ()) thm (fn _ => [Blast_tac 1]);
-
-int_prove "conj_commute" "P&Q <-> Q&P";
-int_prove "conj_left_commute" "P&(Q&R) <-> Q&(P&R)";
-bind_thms ("conj_comms", [conj_commute, conj_left_commute]);
-
-int_prove "disj_commute" "P|Q <-> Q|P";
-int_prove "disj_left_commute" "P|(Q|R) <-> Q|(P|R)";
-bind_thms ("disj_comms", [disj_commute, disj_left_commute]);
-
-int_prove "conj_disj_distribL" "P&(Q|R) <-> (P&Q | P&R)";
-int_prove "conj_disj_distribR" "(P|Q)&R <-> (P&R | Q&R)";
-
-int_prove "disj_conj_distribL" "P|(Q&R) <-> (P|Q) & (P|R)";
-int_prove "disj_conj_distribR" "(P&Q)|R <-> (P|R) & (Q|R)";
-
-int_prove "imp_conj_distrib" "(P --> (Q&R)) <-> (P-->Q) & (P-->R)";
-int_prove "imp_conj"         "((P&Q)-->R)   <-> (P --> (Q --> R))";
-int_prove "imp_disj"         "(P|Q --> R)   <-> (P-->R) & (Q-->R)";
-
-prove "imp_disj1" "(P-->Q) | R <-> (P-->Q | R)";
-prove "imp_disj2" "Q | (P-->R) <-> (P-->Q | R)";
-
-int_prove "de_Morgan_disj" "(~(P | Q)) <-> (~P & ~Q)";
-prove     "de_Morgan_conj" "(~(P & Q)) <-> (~P | ~Q)";
-
-prove     "not_imp" "~(P --> Q) <-> (P & ~Q)";
-prove     "not_iff" "~(P <-> Q) <-> (P <-> ~Q)";
-
-prove     "not_all" "(~ (ALL x. P(x))) <-> (EX x.~P(x))";
-prove     "imp_all" "((ALL x. P(x)) --> Q) <-> (EX x. P(x) --> Q)";
-int_prove "not_ex"  "(~ (EX x. P(x))) <-> (ALL x.~P(x))";
-int_prove "imp_ex" "((EX x. P(x)) --> Q) <-> (ALL x. P(x) --> Q)";
-
-int_prove "ex_disj_distrib"
-    "(EX x. P(x) | Q(x)) <-> ((EX x. P(x)) | (EX x. Q(x)))";
-int_prove "all_conj_distrib"
-    "(ALL x. P(x) & Q(x)) <-> ((ALL x. P(x)) & (ALL x. Q(x)))";
-
 
 (** make simplification procedures for quantifier elimination **)
 structure Quantifier1 = Quantifier1Fun(
@@ -221,20 +69,20 @@
   val conj = FOLogic.conj
   val imp  = FOLogic.imp
   (*rules*)
-  val iff_reflection = iff_reflection
-  val iffI = iffI
-  val iff_trans = iff_trans
-  val conjI= conjI
-  val conjE= conjE
-  val impI = impI
-  val mp   = mp
-  val uncurry = thm "uncurry"
-  val exI  = exI
-  val exE  = exE
-  val iff_allI = thm "iff_allI"
-  val iff_exI = thm "iff_exI"
-  val all_comm = thm "all_comm"
-  val ex_comm = thm "ex_comm"
+  val iff_reflection = @{thm iff_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm iff_trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
 end);
 
 val defEX_regroup =
@@ -252,14 +100,14 @@
   struct
   structure Simplifier = Simplifier
   val mk_eq          = mk_eq
-  val meta_eq_to_iff = meta_eq_to_iff
-  val iffD           = iffD2
-  val disjE          = disjE
-  val conjE          = conjE
-  val exE            = exE
-  val contrapos      = thm "contrapos"
-  val contrapos2     = thm "contrapos2"
-  val notnotD        = thm "notnotD"
+  val meta_eq_to_iff = @{thm meta_eq_to_iff}
+  val iffD           = @{thm iffD2}
+  val disjE          = @{thm disjE}
+  val conjE          = @{thm conjE}
+  val exE            = @{thm exE}
+  val contrapos      = @{thm contrapos}
+  val contrapos2     = @{thm contrapos2}
+  val notnotD        = @{thm notnotD}
   end;
 
 structure Splitter = SplitterFun(SplitterData);
@@ -275,27 +123,17 @@
 
 (*** Standard simpsets ***)
 
-bind_thms ("meta_simps",
- [triv_forall_equality,   (* prunes params *)
-  thm "True_implies_equals"]);  (* prune asms `True' *)
-
-bind_thms ("IFOL_simps",
- [refl RS P_iff_T] @ conj_simps @ disj_simps @ not_simps @
-  imp_simps @ iff_simps @ quant_simps);
+val triv_rls = [@{thm TrueI}, @{thm refl}, reflexive_thm, @{thm iff_refl}, @{thm notFalseI}];
 
-bind_thm ("notFalseI", int_prove_fun "~False");
-bind_thms ("triv_rls",
-  [TrueI, refl, reflexive_thm, iff_refl, thm "notFalseI"]);
-
-fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls@prems),
-                                 atac, etac FalseE];
+fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls @ prems),
+                                 atac, etac @{thm FalseE}];
 (*No premature instantiation of variables during simplification*)
-fun   safe_solver prems = FIRST'[match_tac (triv_rls@prems),
-                                 eq_assume_tac, ematch_tac [FalseE]];
+fun   safe_solver prems = FIRST'[match_tac (triv_rls @ prems),
+                                 eq_assume_tac, ematch_tac [@{thm FalseE}]];
 
 (*No simprules, but basic infastructure for simplification*)
 val FOL_basic_ss =
-  Simplifier.theory_context (the_context ()) empty_ss
+  Simplifier.theory_context @{theory} empty_ss
   setsubgoaler asm_simp_tac
   setSSolver (mk_solver "FOL safe" safe_solver)
   setSolver (mk_solver "FOL unsafe" unsafe_solver)
@@ -310,21 +148,12 @@
 (*intuitionistic simprules only*)
 val IFOL_ss =
   FOL_basic_ss
-  addsimps (meta_simps @ IFOL_simps @ int_ex_simps @ int_all_simps)
+  addsimps (@{thms meta_simps} @ @{thms IFOL_simps} @ @{thms int_ex_simps} @ @{thms int_all_simps})
   addsimprocs [defALL_regroup, defEX_regroup]    
-  addcongs [thm "imp_cong"];
-
-bind_thms ("cla_simps",
-  [de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2,
-   not_imp, not_all, not_ex, cases_simp] @
-  map prove_fun
-   ["~(P&Q) <-> ~P | ~Q",
-    "P | ~P",             "~P | P",
-    "~ ~ P <-> P",        "(~P --> P) <-> P",
-    "(~P <-> ~Q) <-> (P<->Q)"]);
+  addcongs [@{thm imp_cong}];
 
 (*classical simprules too*)
-val FOL_ss = IFOL_ss addsimps (cla_simps @ cla_ex_simps @ cla_all_simps);
+val FOL_ss = IFOL_ss addsimps (@{thms cla_simps} @ @{thms cla_ex_simps} @ @{thms cla_all_simps});
 
 val simpsetup = (fn thy => (change_simpset_of thy (fn _ => FOL_ss); thy));
 
@@ -334,7 +163,7 @@
 structure Clasimp = ClasimpFun
  (structure Simplifier = Simplifier and Splitter = Splitter
   and Classical  = Cla and Blast = Blast
-  val iffD1 = iffD1 val iffD2 = iffD2 val notE = notE);
+  val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val notE = @{thm notE});
 open Clasimp;
 
 ML_Context.value_antiq "clasimpset"