TFL/post.sml
author paulson
Thu, 15 May 1997 12:29:59 +0200
changeset 3191 14bd6e5985f1
parent 2467 357adb429fda
child 3208 8336393de482
permissions -rw-r--r--
TFL now integrated with HOL (more work needed)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
     1
signature TFL = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
     2
  sig
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     3
   structure Prim : TFL_sig
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     4
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
     5
   val tgoalw : theory -> thm list -> thm list -> thm list
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
     6
   val tgoal: theory -> thm list -> thm list
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     7
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     8
   val WF_TAC : thm list -> tactic
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
     9
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    10
   val simplifier : thm -> thm
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    11
   val std_postprocessor : theory 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    12
                           -> {induction:thm, rules:thm, TCs:term list list} 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    13
                           -> {induction:thm, rules:thm, nested_tcs:thm list}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    14
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    15
   val define_i : theory -> term -> term -> theory * (thm * Prim.pattern list)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    16
   val define   : theory -> string -> string list -> theory * Prim.pattern list
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    17
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    18
   val simplify_defn : theory * (string * Prim.pattern list)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    19
                        -> {rules:thm list, induct:thm, tcs:term list}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    20
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    21
  (*-------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    22
       val function : theory -> term -> {theory:theory, eq_ind : thm}
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    23
       val lazyR_def: theory -> term -> {theory:theory, eqns : thm}
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    24
   *-------------------------------------------------------------------------*)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    25
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    26
   val tflcongs : theory -> thm list
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    27
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    28
  end;
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    29
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    30
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    31
structure Tfl: TFL =
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    32
struct
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    33
 structure Prim = Prim
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    34
 structure S = Prim.USyntax
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    35
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    36
(*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    37
 * Extract termination goals so that they can be put it into a goalstack, or 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    38
 * have a tactic directly applied to them.
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    39
 *--------------------------------------------------------------------------*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    40
fun termination_goals rules = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    41
    map (Logic.freeze_vars o S.drop_Trueprop)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    42
      (foldr (fn (th,A) => union_term (prems_of th, A)) (rules, []));
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    43
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    44
 (*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    45
  * Finds the termination conditions in (highly massaged) definition and 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    46
  * puts them into a goalstack.
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    47
  *--------------------------------------------------------------------------*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    48
 fun tgoalw thy defs rules = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    49
    let val L = termination_goals rules
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    50
        open USyntax
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    51
        val g = cterm_of (sign_of thy) (mk_prop(list_mk_conj L))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    52
    in goalw_cterm defs g
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    53
    end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    54
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    55
 val tgoal = Utils.C tgoalw [];
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    56
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    57
 (*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    58
  * Simple wellfoundedness prover.
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    59
  *--------------------------------------------------------------------------*)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    60
 fun WF_TAC thms = REPEAT(FIRST1(map rtac thms))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    61
 val WFtac = WF_TAC[wf_measure, wf_inv_image, wf_lex_prod, 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    62
                    wf_pred_nat, wf_pred_list, wf_trancl];
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    63
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    64
 val terminator = simp_tac(HOL_ss addsimps[pred_nat_def,pred_list_def,
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    65
                                           fst_conv,snd_conv,
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    66
                                           mem_Collect_eq,lessI]) 1
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    67
                  THEN TRY(fast_tac set_cs 1);
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    68
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    69
val length_Cons = prove_goal List.thy "length(h#t) = Suc(length t)" 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    70
    (fn _ => [Simp_tac 1]);
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    71
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    72
 val simpls = [less_eq RS eq_reflection,
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    73
               lex_prod_def, measure_def, inv_image_def, 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    74
               fst_conv RS eq_reflection, snd_conv RS eq_reflection,
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    75
               mem_Collect_eq RS eq_reflection, length_Cons RS eq_reflection];
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    76
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    77
 (*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    78
  * Does some standard things with the termination conditions of a definition:
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    79
  * attempts to prove wellfoundedness of the given relation; simplifies the
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    80
  * non-proven termination conditions; and finally attempts to prove the 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    81
  * simplified termination conditions.
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    82
  *--------------------------------------------------------------------------*)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    83
 val std_postprocessor = Prim.postprocess{WFtac = WFtac,
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    84
                                    terminator = terminator, 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    85
                                    simplifier = Prim.Rules.simpl_conv simpls};
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    86
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    87
 val simplifier = rewrite_rule (simpls @ #simps(rep_ss HOL_ss) @ 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    88
                                [pred_nat_def,pred_list_def]);
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    89
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    90
 fun tflcongs thy = Prim.Context.read() @ (#case_congs(Thry.extract_info thy));
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    91
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    92
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    93
val concl = #2 o Prim.Rules.dest_thm;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    94
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    95
(*---------------------------------------------------------------------------
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    96
 * Defining a function with an associated termination relation. 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    97
 *---------------------------------------------------------------------------*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    98
fun define_i thy R eqs = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    99
  let val dummy = require_thy thy "WF_Rel" "recursive function definitions";
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   100
      
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   101
      val {functional,pats} = Prim.mk_functional thy eqs
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   102
      val (thm,thry) = Prim.wfrec_definition0 thy  R functional
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   103
  in (thry,(thm,pats))
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   104
  end;
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   105
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   106
(*lcp's version: takes strings; doesn't return "thm" 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   107
	(whose signature is a draft and therefore useless) *)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   108
fun define thy R eqs = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   109
  let fun read thy = readtm (sign_of thy) (TVar(("DUMMY",0),[])) 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   110
      val (thy',(_,pats)) =
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   111
	     define_i thy (read thy R) 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   112
	              (fold_bal (app Ind_Syntax.conj) (map (read thy) eqs))
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   113
  in  (thy',pats)  end
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   114
  handle Utils.ERR {mesg,...} => error mesg;
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   115
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   116
(*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   117
 * Postprocess a definition made by "define". This is a separate stage of 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   118
 * processing from the definition stage.
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   119
 *---------------------------------------------------------------------------*)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   120
local 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   121
structure R = Prim.Rules
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   122
structure U = Utils
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   123
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   124
(* The rest of these local definitions are for the tricky nested case *)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   125
val solved = not o U.can S.dest_eq o #2 o S.strip_forall o concl
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   126
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   127
fun id_thm th = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   128
   let val {lhs,rhs} = S.dest_eq(#2(S.strip_forall(#2 (R.dest_thm th))))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   129
   in S.aconv lhs rhs
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   130
   end handle _ => false
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   131
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   132
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   133
val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   134
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   135
fun mk_meta_eq r = case concl_of r of
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   136
     Const("==",_)$_$_ => r
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   137
  |   _$(Const("op =",_)$_$_) => r RS eq_reflection
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   138
  |   _ => r RS P_imp_P_eq_True
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   139
fun rewrite L = rewrite_rule (map mk_meta_eq (Utils.filter(not o id_thm) L))
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   140
fun reducer thl = rewrite (map standard thl @ #simps(rep_ss HOL_ss))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   141
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   142
fun join_assums th = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   143
  let val {sign,...} = rep_thm th
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   144
      val tych = cterm_of sign
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   145
      val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th)))
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   146
      val cntxtl = (#1 o S.strip_imp) lhs  (* cntxtl should = cntxtr *)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   147
      val cntxtr = (#1 o S.strip_imp) rhs  (* but union is solider *)
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   148
      val cntxt = U.union S.aconv cntxtl cntxtr
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   149
  in 
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   150
    R.GEN_ALL 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   151
      (R.DISCH_ALL 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   152
         (rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th)))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   153
  end
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   154
  val gen_all = S.gen_all
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   155
in
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   156
(*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   157
 * The "reducer" argument is 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   158
 *  (fn thl => rewrite (map standard thl @ #simps(rep_ss HOL_ss))); 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   159
 *---------------------------------------------------------------------------*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   160
fun proof_stage theory reducer {f, R, rules, full_pats_TCs, TCs} =
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   161
  let val dummy = output(std_out, "Proving induction theorem..  ")
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   162
      val ind = Prim.mk_induction theory f R full_pats_TCs
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   163
      val dummy = output(std_out, "Proved induction theorem.\n")
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   164
      val pp = std_postprocessor theory
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   165
      val dummy = output(std_out, "Postprocessing..  ")
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   166
      val {rules,induction,nested_tcs} = pp{rules=rules,induction=ind,TCs=TCs}
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   167
  in
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   168
  case nested_tcs
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   169
  of [] => (output(std_out, "Postprocessing done.\n");
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   170
            {induction=induction, rules=rules,tcs=[]})
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   171
  | L  => let val dummy = output(std_out, "Simplifying nested TCs..  ")
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   172
              val (solved,simplified,stubborn) =
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   173
               U.itlist (fn th => fn (So,Si,St) =>
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   174
                     if (id_thm th) then (So, Si, th::St) else
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   175
                     if (solved th) then (th::So, Si, St) 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   176
                     else (So, th::Si, St)) nested_tcs ([],[],[])
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   177
              val simplified' = map join_assums simplified
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   178
              val induction' = reducer (solved@simplified') induction
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   179
              val rules' = reducer (solved@simplified') rules
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   180
              val dummy = output(std_out, "Postprocessing done.\n")
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   181
          in
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   182
          {induction = induction',
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   183
               rules = rules',
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   184
                 tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   185
                           (simplified@stubborn)}
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   186
          end
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   187
  end handle (e as Utils.ERR _) => Utils.Raise e
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   188
          |   e                 => print_exn e;
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   189
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   190
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   191
(*lcp: put a theorem into Isabelle form, using meta-level connectives*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   192
val meta_outer = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   193
    standard o rule_by_tactic (REPEAT_FIRST (resolve_tac [allI, impI, conjI]));
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   194
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   195
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   196
(*Strip off the outer !P*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   197
val spec'= read_instantiate [("x","P::?'b=>bool")] spec;
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   198
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   199
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   200
fun simplify_defn (thy,(id,pats)) =
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   201
   let val def = freezeT(get_def thy id  RS  meta_eq_to_obj_eq)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   202
       val {theory,rules,TCs,full_pats_TCs,patterns} = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   203
		Prim.post_definition (thy,(def,pats))
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   204
       val {lhs=f,rhs} = S.dest_eq(concl def)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   205
       val (_,[R,_]) = S.strip_comb rhs
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   206
       val {induction, rules, tcs} = 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   207
             proof_stage theory reducer
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   208
	       {f = f, R = R, rules = rules,
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   209
		full_pats_TCs = full_pats_TCs,
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   210
		TCs = TCs}
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   211
       val rules' = map (standard o normalize_thm [RSmp]) (R.CONJUNCTS rules)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   212
   in  {induct = meta_outer
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   213
	          (normalize_thm [RSspec,RSmp] (induction RS spec')), 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   214
	rules = rules', 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   215
	tcs = (termination_goals rules') @ tcs}
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   216
   end
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   217
  handle Utils.ERR {mesg,...} => error mesg
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   218
end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   219
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   220
(*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   221
 *
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   222
 *     Definitions with synthesized termination relation temporarily
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   223
 *     deleted -- it's not clear how to integrate this facility with
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   224
 *     the Isabelle theory file scheme, which restricts
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   225
 *     inference at theory-construction time.
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   226
 *
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   227
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   228
local fun floutput s = (output(std_out,s); flush_out std_out)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   229
      structure R = Prim.Rules
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   230
in
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   231
fun function theory eqs = 
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   232
 let val dummy = floutput "Making definition..   "
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   233
     val {rules,R,theory,full_pats_TCs,...} = Prim.lazyR_def theory eqs
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   234
     val f = func_of_cond_eqn (concl(R.CONJUNCT1 rules handle _ => rules))
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   235
     val dummy = floutput "Definition made.\n"
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   236
     val dummy = floutput "Proving induction theorem..  "
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   237
     val induction = Prim.mk_induction theory f R full_pats_TCs
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   238
     val dummy = floutput "Induction theorem proved.\n"
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   239
 in {theory = theory, 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   240
     eq_ind = standard (induction RS (rules RS conjI))}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   241
 end
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   242
 handle (e as Utils.ERR _) => Utils.Raise e
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   243
      |     e              => print_exn e
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   244
end;
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   245
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   246
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   247
fun lazyR_def theory eqs = 
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   248
   let val {rules,theory, ...} = Prim.lazyR_def theory eqs
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   249
   in {eqns=rules, theory=theory}
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   250
   end
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   251
   handle (e as Utils.ERR _) => Utils.Raise e
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   252
        |     e              => print_exn e;
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   253
 *
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   254
 *
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   255
 *---------------------------------------------------------------------------*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   256
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   257
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   258
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   259
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   260
(*---------------------------------------------------------------------------
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   261
 * Install the basic context notions. Others (for nat and list and prod) 
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   262
 * have already been added in thry.sml
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   263
 *---------------------------------------------------------------------------*)
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   264
val () = Prim.Context.write[Thms.LET_CONG, Thms.COND_CONG];
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   265
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   266
end;