diff -r cf19faf11bbd -r 34b2c1bb7178 src/HOL/cladata.ML --- a/src/HOL/cladata.ML Tue Oct 10 10:34:43 2006 +0200 +++ b/src/HOL/cladata.ML Tue Oct 10 10:35:24 2006 +0200 @@ -6,58 +6,53 @@ Setting up the classical reasoner. *) - -(** Applying HypsubstFun to generate hyp_subst_tac **) -section "Classical Reasoner"; - structure Hypsubst_Data = - struct +struct structure Simplifier = Simplifier - (*Take apart an equality judgement; otherwise raise Match!*) - fun dest_eq (Const("op =",T) $ t $ u) = (t, u, domain_type T) + fun dest_eq (Const ("op =", T) $ t $ u) = (t, u, domain_type T) val dest_Trueprop = HOLogic.dest_Trueprop val dest_imp = HOLogic.dest_imp - val eq_reflection = eq_reflection - val rev_eq_reflection = def_imp_eq - val imp_intr = impI - val rev_mp = rev_mp - val subst = subst - val sym = sym - val thin_refl = prove_goal (the_context ()) - "!!X. [|x=x; PROP W|] ==> PROP W" (K [atac 1]); - end; + val eq_reflection = HOL.eq_reflection + val rev_eq_reflection = HOL.def_imp_eq + val imp_intr = HOL.impI + val rev_mp = HOL.rev_mp + val subst = HOL.subst + val sym = HOL.sym + val thin_refl = thm "thin_refl"; +end; structure Hypsubst = HypsubstFun(Hypsubst_Data); + +structure Classical_Data = +struct + val mp = HOL.mp + val not_elim = HOL.notE + val classical = HOL.classical + val sizef = Drule.size_of_thm + val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] +end; + +structure Classical = ClassicalFun(Classical_Data); +structure BasicClassical: BASIC_CLASSICAL = Classical; + +structure HOL = +struct + +open HOL; open Hypsubst; +open BasicClassical; (*prevent substitution on bool*) fun hyp_subst_tac' i thm = if i <= Thm.nprems_of thm andalso Term.exists_Const (fn ("op =", Type (_, [T, _])) => T <> Type ("bool", []) | _ => false) - (List.nth (Thm.prems_of thm, i - 1)) then hyp_subst_tac i thm else no_tac thm; - - -(*** Applying ClassicalFun to create a classical prover ***) -structure Classical_Data = - struct - val mp = mp - val not_elim = notE - val classical = classical - val sizef = size_of_thm - val hyp_subst_tacs=[hyp_subst_tac] - end; - -structure Classical = ClassicalFun(Classical_Data); - -structure BasicClassical: BASIC_CLASSICAL = Classical; - -open BasicClassical; + (nth (Thm.prems_of thm) (i - 1)) then hyp_subst_tac i thm else no_tac thm; (*Propositional rules*) -val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI] - addSEs [conjE,disjE,impCE,FalseE,iffCE]; +val prop_cs = empty_cs addSIs [HOL.refl, HOL.TrueI, HOL.conjI, HOL.disjCI, HOL.impI, HOL.notI, HOL.iffI] + addSEs [HOL.conjE, HOL.disjE, HOL.impCE, HOL.FalseE, HOL.iffCE]; (*Quantifier rules*) -val HOL_cs = prop_cs addSIs [allI,ex_ex1I] addIs [exI, the_equality] - addSEs [exE] addEs [allE]; +val cs = prop_cs addSIs [HOL.allI, HOL.ex_ex1I] addIs [HOL.exI, HOL.the_equality] + addSEs [HOL.exE] addEs [HOL.allE]; -val clasetup = (fn thy => (change_claset_of thy (fn _ => HOL_cs); thy)); +end;