src/HOL/Tools/TFL/post.ML
changeset 59582 0fbed69ff081
parent 59580 cbc38731d42f
child 59621 291934bac95e
equal deleted inserted replaced
59580:cbc38731d42f 59582:0fbed69ff081
    24  * Extract termination goals so that they can be put it into a goalstack, or
    24  * Extract termination goals so that they can be put it into a goalstack, or
    25  * have a tactic directly applied to them.
    25  * have a tactic directly applied to them.
    26  *--------------------------------------------------------------------------*)
    26  *--------------------------------------------------------------------------*)
    27 fun termination_goals rules =
    27 fun termination_goals rules =
    28     map (Type.legacy_freeze o HOLogic.dest_Trueprop)
    28     map (Type.legacy_freeze o HOLogic.dest_Trueprop)
    29       (fold_rev (union (op aconv) o prems_of) rules []);
    29       (fold_rev (union (op aconv) o Thm.prems_of) rules []);
    30 
    30 
    31 (*---------------------------------------------------------------------------
    31 (*---------------------------------------------------------------------------
    32  * Three postprocessors are applied to the definition.  It
    32  * Three postprocessors are applied to the definition.  It
    33  * attempts to prove wellfoundedness of the given relation, simplifies the
    33  * attempts to prove wellfoundedness of the given relation, simplifies the
    34  * non-proved termination conditions, and finally attempts to prove the
    34  * non-proved termination conditions, and finally attempts to prove the
    60    let val {lhs,rhs} = USyntax.dest_eq (#2 (USyntax.strip_forall (#2 (Rules.dest_thm th))));
    60    let val {lhs,rhs} = USyntax.dest_eq (#2 (USyntax.strip_forall (#2 (Rules.dest_thm th))));
    61    in lhs aconv rhs end
    61    in lhs aconv rhs end
    62    handle Utils.ERR _ => false;
    62    handle Utils.ERR _ => false;
    63    
    63    
    64 val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
    64 val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection;
    65 fun mk_meta_eq r = case concl_of r of
    65 fun mk_meta_eq r =
       
    66   (case Thm.concl_of r of
    66      Const(@{const_name Pure.eq},_)$_$_ => r
    67      Const(@{const_name Pure.eq},_)$_$_ => r
    67   |   _ $(Const(@{const_name HOL.eq},_)$_$_) => r RS eq_reflection
    68   |   _ $(Const(@{const_name HOL.eq},_)$_$_) => r RS eq_reflection
    68   |   _ => r RS P_imp_P_eq_True
    69   |   _ => r RS P_imp_P_eq_True)
    69 
    70 
    70 (*Is this the best way to invoke the simplifier??*)
    71 (*Is this the best way to invoke the simplifier??*)
    71 fun rewrite ctxt L = rewrite_rule ctxt (map mk_meta_eq (filter_out id_thm L))
    72 fun rewrite ctxt L = rewrite_rule ctxt (map mk_meta_eq (filter_out id_thm L))
    72 
    73 
    73 fun join_assums ctxt th =
    74 fun join_assums ctxt th =
    74   let val thy = Thm.theory_of_thm th
    75   let val thy = Thm.theory_of_thm th
    75       val tych = cterm_of thy
    76       val tych = Thm.cterm_of thy
    76       val {lhs,rhs} = USyntax.dest_eq(#2 (USyntax.strip_forall (concl th)))
    77       val {lhs,rhs} = USyntax.dest_eq(#2 (USyntax.strip_forall (concl th)))
    77       val cntxtl = (#1 o USyntax.strip_imp) lhs  (* cntxtl should = cntxtr *)
    78       val cntxtl = (#1 o USyntax.strip_imp) lhs  (* cntxtl should = cntxtr *)
    78       val cntxtr = (#1 o USyntax.strip_imp) rhs  (* but union is solider *)
    79       val cntxtr = (#1 o USyntax.strip_imp) rhs  (* but union is solider *)
    79       val cntxt = union (op aconv) cntxtl cntxtr
    80       val cntxt = union (op aconv) cntxtl cntxtr
    80   in
    81   in
   118   end;
   119   end;
   119 
   120 
   120 
   121 
   121 (*lcp: curry the predicate of the induction rule*)
   122 (*lcp: curry the predicate of the induction rule*)
   122 fun curry_rule ctxt rl =
   123 fun curry_rule ctxt rl =
   123   Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl;
   124   Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (Thm.concl_of rl))) rl;
   124 
   125 
   125 (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
   126 (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
   126 fun meta_outer ctxt =
   127 fun meta_outer ctxt =
   127   curry_rule ctxt o Drule.export_without_context o
   128   curry_rule ctxt o Drule.export_without_context o
   128   rule_by_tactic ctxt (REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' etac conjE)));
   129   rule_by_tactic ctxt (REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' etac conjE)));
   185  *---------------------------------------------------------------------------*)
   186  *---------------------------------------------------------------------------*)
   186 fun define_i strict ctxt congs wfs fid R eqs thy =
   187 fun define_i strict ctxt congs wfs fid R eqs thy =
   187   let val {functional,pats} = Prim.mk_functional thy eqs
   188   let val {functional,pats} = Prim.mk_functional thy eqs
   188       val (thy, def) = Prim.wfrec_definition0 thy fid R functional
   189       val (thy, def) = Prim.wfrec_definition0 thy fid R functional
   189       val ctxt = Proof_Context.transfer thy ctxt
   190       val ctxt = Proof_Context.transfer thy ctxt
   190       val (lhs, _) = Logic.dest_equals (prop_of def)
   191       val (lhs, _) = Logic.dest_equals (Thm.prop_of def)
   191       val {induct, rules, tcs} = simplify_defn strict thy ctxt congs wfs fid pats def
   192       val {induct, rules, tcs} = simplify_defn strict thy ctxt congs wfs fid pats def
   192       val rules' = 
   193       val rules' = 
   193           if strict then derive_init_eqs ctxt rules eqs
   194           if strict then derive_init_eqs ctxt rules eqs
   194           else rules
   195           else rules
   195   in ({lhs = lhs, rules = rules', induct = induct, tcs = tcs}, thy) end;
   196   in ({lhs = lhs, rules = rules', induct = induct, tcs = tcs}, thy) end;