wenzelm@10769: (* Title: TFL/post.ML wenzelm@10769: ID: $Id$ wenzelm@10769: Author: Konrad Slind, Cambridge University Computer Laboratory wenzelm@10769: Copyright 1997 University of Cambridge wenzelm@10769: wenzelm@10769: Second part of main module (postprocessing of TFL definitions). wenzelm@10769: *) wenzelm@10769: wenzelm@10769: signature TFL = wenzelm@10769: sig wenzelm@10769: val trace: bool ref wenzelm@10769: val quiet_mode: bool ref wenzelm@10769: val message: string -> unit wenzelm@10769: val tgoalw: theory -> thm list -> thm list -> thm list wenzelm@10769: val tgoal: theory -> thm list -> thm list wenzelm@11632: val define_i: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring -> wenzelm@10769: term -> term list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list} wenzelm@11632: val define: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring -> wenzelm@10769: string -> string list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list} wenzelm@10769: val defer_i: theory -> thm list -> xstring -> term list -> theory * thm wenzelm@10769: val defer: theory -> thm list -> xstring -> string list -> theory * thm wenzelm@10769: end; wenzelm@10769: wenzelm@10769: structure Tfl: TFL = wenzelm@10769: struct wenzelm@10769: wenzelm@10769: structure S = USyntax wenzelm@10769: wenzelm@10769: wenzelm@10769: (* messages *) wenzelm@10769: wenzelm@10769: val trace = Prim.trace wenzelm@10769: wenzelm@10769: val quiet_mode = ref false; wenzelm@10769: fun message s = if ! quiet_mode then () else writeln s; wenzelm@10769: wenzelm@10769: wenzelm@10769: (* misc *) wenzelm@10769: wenzelm@10769: (*--------------------------------------------------------------------------- wenzelm@10769: * Extract termination goals so that they can be put it into a goalstack, or wenzelm@10769: * have a tactic directly applied to them. wenzelm@10769: *--------------------------------------------------------------------------*) wenzelm@10769: fun termination_goals rules = wenzelm@16287: map (Type.freeze o HOLogic.dest_Trueprop) wenzelm@18139: (foldr (fn (th,A) => gen_union (op aconv) (prems_of th, A)) [] rules); wenzelm@10769: wenzelm@10769: (*--------------------------------------------------------------------------- wenzelm@10769: * Finds the termination conditions in (highly massaged) definition and wenzelm@10769: * puts them into a goalstack. wenzelm@10769: *--------------------------------------------------------------------------*) wenzelm@10769: fun tgoalw thy defs rules = wenzelm@10769: case termination_goals rules of wenzelm@10769: [] => error "tgoalw: no termination conditions to prove" wenzelm@17959: | L => OldGoals.goalw_cterm defs wenzelm@10769: (Thm.cterm_of (Theory.sign_of thy) wenzelm@10769: (HOLogic.mk_Trueprop(USyntax.list_mk_conj L))); wenzelm@10769: wenzelm@10769: fun tgoal thy = tgoalw thy []; wenzelm@10769: wenzelm@10769: (*--------------------------------------------------------------------------- wenzelm@10769: * Three postprocessors are applied to the definition. It wenzelm@10769: * attempts to prove wellfoundedness of the given relation, simplifies the wenzelm@10769: * non-proved termination conditions, and finally attempts to prove the wenzelm@10769: * simplified termination conditions. wenzelm@10769: *--------------------------------------------------------------------------*) wenzelm@11632: fun std_postprocessor strict cs ss wfs = wenzelm@11632: Prim.postprocess strict wenzelm@10769: {wf_tac = REPEAT (ares_tac wfs 1), wenzelm@10769: terminator = asm_simp_tac ss 1 nipkow@13501: THEN TRY (silent_arith_tac 1 ORELSE nipkow@12488: fast_tac (cs addSDs [not0_implies_Suc] addss ss) 1), wenzelm@10769: simplifier = Rules.simpl_conv ss []}; wenzelm@10769: wenzelm@10769: wenzelm@10769: wenzelm@10769: val concl = #2 o Rules.dest_thm; wenzelm@10769: wenzelm@10769: (*--------------------------------------------------------------------------- wenzelm@10769: * Postprocess a definition made by "define". This is a separate stage of wenzelm@10769: * processing from the definition stage. wenzelm@10769: *---------------------------------------------------------------------------*) wenzelm@10769: local wenzelm@10769: structure R = Rules wenzelm@10769: structure U = Utils wenzelm@10769: wenzelm@10769: (* The rest of these local definitions are for the tricky nested case *) wenzelm@10769: val solved = not o can S.dest_eq o #2 o S.strip_forall o concl wenzelm@10769: wenzelm@10769: fun id_thm th = wenzelm@10769: let val {lhs,rhs} = S.dest_eq (#2 (S.strip_forall (#2 (R.dest_thm th)))); wenzelm@10769: in lhs aconv rhs end wenzelm@10769: handle U.ERR _ => false; wenzelm@10769: wenzelm@10769: wenzelm@10769: fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]); wenzelm@10769: val P_imp_P_iff_True = prover "P --> (P= True)" RS mp; wenzelm@10769: val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection; wenzelm@10769: fun mk_meta_eq r = case concl_of r of wenzelm@10769: Const("==",_)$_$_ => r wenzelm@10769: | _ $(Const("op =",_)$_$_) => r RS eq_reflection wenzelm@10769: | _ => r RS P_imp_P_eq_True wenzelm@10769: wenzelm@10769: (*Is this the best way to invoke the simplifier??*) skalberg@15570: fun rewrite L = rewrite_rule (map mk_meta_eq (List.filter(not o id_thm) L)) wenzelm@10769: wenzelm@10769: fun join_assums th = wenzelm@10769: let val {sign,...} = rep_thm th wenzelm@10769: val tych = cterm_of sign wenzelm@10769: val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th))) wenzelm@10769: val cntxtl = (#1 o S.strip_imp) lhs (* cntxtl should = cntxtr *) wenzelm@10769: val cntxtr = (#1 o S.strip_imp) rhs (* but union is solider *) wenzelm@10769: val cntxt = gen_union (op aconv) (cntxtl, cntxtr) wenzelm@10769: in wenzelm@10769: R.GEN_ALL wenzelm@10769: (R.DISCH_ALL wenzelm@10769: (rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th))) wenzelm@10769: end wenzelm@10769: val gen_all = S.gen_all wenzelm@10769: in wenzelm@11632: fun proof_stage strict cs ss wfs theory {f, R, rules, full_pats_TCs, TCs} = wenzelm@10769: let wenzelm@10769: val _ = message "Proving induction theorem ..." wenzelm@10769: val ind = Prim.mk_induction theory {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs} wenzelm@10769: val _ = message "Postprocessing ..."; wenzelm@10769: val {rules, induction, nested_tcs} = wenzelm@11632: std_postprocessor strict cs ss wfs theory {rules=rules, induction=ind, TCs=TCs} wenzelm@10769: in wenzelm@10769: case nested_tcs wenzelm@10769: of [] => {induction=induction, rules=rules,tcs=[]} wenzelm@10769: | L => let val dummy = message "Simplifying nested TCs ..." wenzelm@10769: val (solved,simplified,stubborn) = wenzelm@16852: fold_rev (fn th => fn (So,Si,St) => wenzelm@10769: if (id_thm th) then (So, Si, th::St) else wenzelm@10769: if (solved th) then (th::So, Si, St) wenzelm@10769: else (So, th::Si, St)) nested_tcs ([],[],[]) wenzelm@10769: val simplified' = map join_assums simplified paulson@14240: val dummy = (Prim.trace_thms "solved =" solved; paulson@14240: Prim.trace_thms "simplified' =" simplified') wenzelm@10769: val rewr = full_simplify (ss addsimps (solved @ simplified')); paulson@14240: val dummy = Prim.trace_thms "Simplifying the induction rule..." paulson@14240: [induction] wenzelm@10769: val induction' = rewr induction paulson@14240: val dummy = Prim.trace_thms "Simplifying the recursion rules..." paulson@14240: [rules] paulson@14240: val rules' = rewr rules paulson@14240: val _ = message "... Postprocessing finished"; wenzelm@10769: in wenzelm@10769: {induction = induction', wenzelm@10769: rules = rules', wenzelm@10769: tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl) wenzelm@10769: (simplified@stubborn)} wenzelm@10769: end wenzelm@10769: end; wenzelm@10769: wenzelm@10769: wenzelm@10769: (*lcp: curry the predicate of the induction rule*) wenzelm@11038: fun curry_rule rl = wenzelm@19736: SplitRule.split_rule_var (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl; wenzelm@10769: wenzelm@10769: (*lcp: put a theorem into Isabelle form, using meta-level connectives*) wenzelm@10769: val meta_outer = wenzelm@11038: curry_rule o standard o wenzelm@11038: rule_by_tactic (REPEAT (FIRSTGOAL (resolve_tac [allI, impI, conjI] ORELSE' etac conjE))); wenzelm@10769: wenzelm@10769: (*Strip off the outer !P*) wenzelm@10769: val spec'= read_instantiate [("x","P::?'b=>bool")] spec; wenzelm@10769: paulson@14240: fun tracing true _ = () paulson@14240: | tracing false msg = writeln msg; paulson@14240: wenzelm@11632: fun simplify_defn strict thy cs ss congs wfs id pats def0 = wenzelm@19927: let val def = Thm.freezeT def0 RS meta_eq_to_obj_eq wenzelm@20061: val {rules,rows,TCs,full_pats_TCs} = paulson@14240: Prim.post_definition congs (thy, (def,pats)) wenzelm@10769: val {lhs=f,rhs} = S.dest_eq (concl def) wenzelm@10769: val (_,[R,_]) = S.strip_comb rhs paulson@14240: val dummy = Prim.trace_thms "congs =" congs paulson@14240: (*the next step has caused simplifier looping in some cases*) wenzelm@10769: val {induction, rules, tcs} = wenzelm@20061: proof_stage strict cs ss wfs thy wenzelm@10769: {f = f, R = R, rules = rules, wenzelm@10769: full_pats_TCs = full_pats_TCs, wenzelm@10769: TCs = TCs} paulson@14240: val rules' = map (standard o ObjectLogic.rulify_no_asm) paulson@14240: (R.CONJUNCTS rules) paulson@14240: in {induct = meta_outer (ObjectLogic.rulify_no_asm (induction RS spec')), wenzelm@10769: rules = ListPair.zip(rules', rows), wenzelm@10769: tcs = (termination_goals rules') @ tcs} wenzelm@10769: end wenzelm@10769: handle U.ERR {mesg,func,module} => wenzelm@10769: error (mesg ^ wenzelm@10769: "\n (In TFL function " ^ module ^ "." ^ func ^ ")"); wenzelm@10769: paulson@15150: paulson@15150: (* Derive the initial equations from the case-split rules to meet the paulson@15150: users specification of the recursive function. paulson@15150: Note: We don't do this if the wf conditions fail to be solved, as each paulson@15150: case may have a different wf condition. We could group the conditions paulson@15150: together and say that they must be true to solve the general case, paulson@15150: but that would hide from the user which sub-case they were related paulson@15150: to. Probably this is not important, and it would work fine, but, for now, I dixon@15171: prefer leaving more fine-grain control to the user. dixon@15171: -- Lucas Dixon, Aug 2004 *) paulson@15150: local paulson@15150: fun get_related_thms i = skalberg@15570: List.mapPartial ((fn (r,x) => if x = i then SOME r else NONE)); paulson@15150: paulson@15150: fun solve_eq (th, [], i) = wenzelm@18678: error "derive_init_eqs: missing rules" paulson@15150: | solve_eq (th, [a], i) = [(a, i)] paulson@15150: | solve_eq (th, splitths as (_ :: _), i) = dixon@15171: (writeln "Proving unsplit equation..."; paulson@15150: [((standard o ObjectLogic.rulify_no_asm) dixon@15171: (CaseSplit.splitto splitths th), i)]) paulson@15150: (* if there's an error, pretend nothing happened with this definition paulson@15150: We should probably print something out so that the user knows...? *) wenzelm@18678: handle ERROR s => wenzelm@17615: (warning ("recdef (solve_eq): " ^ s); map (fn x => (x,i)) splitths); paulson@15150: in paulson@15150: fun derive_init_eqs sgn rules eqs = paulson@15150: let paulson@15150: val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o HOLogic.mk_Trueprop) paulson@15150: eqs paulson@15150: fun countlist l = skalberg@15570: (rev o snd o (Library.foldl (fn ((i,L), e) => (i + 1,(e,i) :: L)))) ((0,[]), l) paulson@15150: in skalberg@15570: List.concat (map (fn (e,i) => solve_eq (e, (get_related_thms i rules), i)) paulson@15150: (countlist eqths)) paulson@15150: end; paulson@15150: end; paulson@15150: paulson@15150: wenzelm@10769: (*--------------------------------------------------------------------------- wenzelm@10769: * Defining a function with an associated termination relation. wenzelm@10769: *---------------------------------------------------------------------------*) wenzelm@11632: fun define_i strict thy cs ss congs wfs fid R eqs = wenzelm@10769: let val {functional,pats} = Prim.mk_functional thy eqs wenzelm@10769: val (thy, def) = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional dixon@15171: val {induct, rules, tcs} = dixon@15171: simplify_defn strict thy cs ss congs wfs fid pats def dixon@15171: val rules' = dixon@15171: if strict then derive_init_eqs (Theory.sign_of thy) rules eqs dixon@15171: else rules dixon@15171: in (thy, {rules = rules', induct = induct, tcs = tcs}) end; wenzelm@10769: wenzelm@11632: fun define strict thy cs ss congs wfs fid R seqs = wenzelm@16975: define_i strict thy cs ss congs wfs fid (Sign.read_term thy R) (map (Sign.read_term thy) seqs) wenzelm@10769: handle U.ERR {mesg,...} => error mesg; wenzelm@10769: wenzelm@10769: wenzelm@10769: (*--------------------------------------------------------------------------- wenzelm@10769: * wenzelm@10769: * Definitions with synthesized termination relation wenzelm@10769: * wenzelm@10769: *---------------------------------------------------------------------------*) wenzelm@10769: wenzelm@10769: fun func_of_cond_eqn tm = wenzelm@10769: #1 (S.strip_comb (#lhs (S.dest_eq (#2 (S.strip_forall (#2 (S.strip_imp tm))))))); wenzelm@10769: wenzelm@10769: fun defer_i thy congs fid eqs = wenzelm@10769: let val {rules,R,theory,full_pats_TCs,SV,...} = wenzelm@10769: Prim.lazyR_def thy (Sign.base_name fid) congs eqs wenzelm@10769: val f = func_of_cond_eqn (concl (R.CONJUNCT1 rules handle U.ERR _ => rules)); wenzelm@10769: val dummy = message "Proving induction theorem ..."; wenzelm@10769: val induction = Prim.mk_induction theory wenzelm@10769: {fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs} wenzelm@10769: in (theory, wenzelm@10769: (*return the conjoined induction rule and recursion equations, wenzelm@10769: with assumptions remaining to discharge*) wenzelm@10769: standard (induction RS (rules RS conjI))) wenzelm@10769: end wenzelm@10769: wenzelm@10769: fun defer thy congs fid seqs = wenzelm@16975: defer_i thy congs fid (map (Sign.read_term thy) seqs) wenzelm@10769: handle U.ERR {mesg,...} => error mesg; wenzelm@10769: end; wenzelm@10769: wenzelm@10769: end;