src/HOL/Tools/TFL/post.ML
author haftmann
Fri, 20 Jul 2007 14:28:05 +0200
changeset 23880 64b9806e160b
parent 23150 073a65f0bc40
child 24075 366d4d234814
permissions -rw-r--r--
dropped Nat.ML legacy bindings
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Tools/TFL/post.ML
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     3
    Author:     Konrad Slind, Cambridge University Computer Laboratory
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     4
    Copyright   1997  University of Cambridge
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     5
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     6
Second part of main module (postprocessing of TFL definitions).
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     7
*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     8
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     9
signature TFL =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    10
sig
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    11
  val trace: bool ref
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    12
  val quiet_mode: bool ref
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    13
  val message: string -> unit
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    14
  val tgoalw: theory -> thm list -> thm list -> thm list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    15
  val tgoal: theory -> thm list -> thm list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    16
  val define_i: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring ->
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    17
    term -> term list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    18
  val define: bool -> theory -> claset -> simpset -> thm list -> thm list -> xstring ->
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    19
    string -> string list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    20
  val defer_i: theory -> thm list -> xstring -> term list -> theory * thm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    21
  val defer: theory -> thm list -> xstring -> string list -> theory * thm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    22
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    23
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    24
structure Tfl: TFL =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    25
struct
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    26
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    27
structure S = USyntax
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    28
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    29
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    30
(* messages *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    31
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    32
val trace = Prim.trace
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    33
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    34
val quiet_mode = ref false;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    35
fun message s = if ! quiet_mode then () else writeln s;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    36
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    37
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    38
(* misc *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    39
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    40
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    41
 * Extract termination goals so that they can be put it into a goalstack, or
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    42
 * have a tactic directly applied to them.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    43
 *--------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    44
fun termination_goals rules =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    45
    map (Type.freeze o HOLogic.dest_Trueprop)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    46
      (foldr (fn (th,A) => gen_union (op aconv) (prems_of th, A)) [] rules);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    47
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    48
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    49
 * Finds the termination conditions in (highly massaged) definition and
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    50
 * puts them into a goalstack.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    51
 *--------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    52
fun tgoalw thy defs rules =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    53
  case termination_goals rules of
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    54
      [] => error "tgoalw: no termination conditions to prove"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    55
    | L  => OldGoals.goalw_cterm defs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    56
              (Thm.cterm_of thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    57
                        (HOLogic.mk_Trueprop(USyntax.list_mk_conj L)));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    58
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    59
fun tgoal thy = tgoalw thy [];
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    60
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    61
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    62
 * Three postprocessors are applied to the definition.  It
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    63
 * attempts to prove wellfoundedness of the given relation, simplifies the
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    64
 * non-proved termination conditions, and finally attempts to prove the
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    65
 * simplified termination conditions.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    66
 *--------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    67
fun std_postprocessor strict cs ss wfs =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    68
  Prim.postprocess strict
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    69
   {wf_tac     = REPEAT (ares_tac wfs 1),
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    70
    terminator = asm_simp_tac ss 1
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    71
                 THEN TRY (silent_arith_tac 1 ORELSE
23880
64b9806e160b dropped Nat.ML legacy bindings
haftmann
parents: 23150
diff changeset
    72
                           fast_tac (cs addSDs [@{thm not0_implies_Suc}] addss ss) 1),
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    73
    simplifier = Rules.simpl_conv ss []};
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    74
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    75
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    76
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    77
val concl = #2 o Rules.dest_thm;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    78
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    79
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    80
 * Postprocess a definition made by "define". This is a separate stage of
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    81
 * processing from the definition stage.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    82
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    83
local
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    84
structure R = Rules
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    85
structure U = Utils
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    86
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    87
(* The rest of these local definitions are for the tricky nested case *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    88
val solved = not o can S.dest_eq o #2 o S.strip_forall o concl
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    89
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    90
fun id_thm th =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    91
   let val {lhs,rhs} = S.dest_eq (#2 (S.strip_forall (#2 (R.dest_thm th))));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    92
   in lhs aconv rhs end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    93
   handle U.ERR _ => false;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    94
   
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    95
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    96
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    97
val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    98
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    99
fun mk_meta_eq r = case concl_of r of
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   100
     Const("==",_)$_$_ => r
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   101
  |   _ $(Const("op =",_)$_$_) => r RS eq_reflection
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   102
  |   _ => r RS P_imp_P_eq_True
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   103
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   104
(*Is this the best way to invoke the simplifier??*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   105
fun rewrite L = rewrite_rule (map mk_meta_eq (List.filter(not o id_thm) L))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   106
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   107
fun join_assums th =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   108
  let val {thy,...} = rep_thm th
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   109
      val tych = cterm_of thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   110
      val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th)))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   111
      val cntxtl = (#1 o S.strip_imp) lhs  (* cntxtl should = cntxtr *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   112
      val cntxtr = (#1 o S.strip_imp) rhs  (* but union is solider *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   113
      val cntxt = gen_union (op aconv) (cntxtl, cntxtr)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   114
  in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   115
    R.GEN_ALL
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   116
      (R.DISCH_ALL
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   117
         (rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th)))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   118
  end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   119
  val gen_all = S.gen_all
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   120
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   121
fun proof_stage strict cs ss wfs theory {f, R, rules, full_pats_TCs, TCs} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   122
  let
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   123
    val _ = message "Proving induction theorem ..."
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   124
    val ind = Prim.mk_induction theory {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   125
    val _ = message "Postprocessing ...";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   126
    val {rules, induction, nested_tcs} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   127
      std_postprocessor strict cs ss wfs theory {rules=rules, induction=ind, TCs=TCs}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   128
  in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   129
  case nested_tcs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   130
  of [] => {induction=induction, rules=rules,tcs=[]}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   131
  | L  => let val dummy = message "Simplifying nested TCs ..."
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   132
              val (solved,simplified,stubborn) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   133
               fold_rev (fn th => fn (So,Si,St) =>
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   134
                     if (id_thm th) then (So, Si, th::St) else
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   135
                     if (solved th) then (th::So, Si, St)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   136
                     else (So, th::Si, St)) nested_tcs ([],[],[])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   137
              val simplified' = map join_assums simplified
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   138
              val dummy = (Prim.trace_thms "solved =" solved;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   139
                           Prim.trace_thms "simplified' =" simplified')
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   140
              val rewr = full_simplify (ss addsimps (solved @ simplified'));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   141
              val dummy = Prim.trace_thms "Simplifying the induction rule..."
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   142
                                          [induction]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   143
              val induction' = rewr induction
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   144
              val dummy = Prim.trace_thms "Simplifying the recursion rules..."
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   145
                                          [rules]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   146
              val rules'     = rewr rules
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   147
              val _ = message "... Postprocessing finished";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   148
          in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   149
          {induction = induction',
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   150
               rules = rules',
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   151
                 tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   152
                           (simplified@stubborn)}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   153
          end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   154
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   155
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   156
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   157
(*lcp: curry the predicate of the induction rule*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   158
fun curry_rule rl =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   159
  SplitRule.split_rule_var (Term.head_of (HOLogic.dest_Trueprop (concl_of rl))) rl;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   160
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   161
(*lcp: put a theorem into Isabelle form, using meta-level connectives*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   162
val meta_outer =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   163
  curry_rule o standard o
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   164
  rule_by_tactic (REPEAT (FIRSTGOAL (resolve_tac [allI, impI, conjI] ORELSE' etac conjE)));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   165
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   166
(*Strip off the outer !P*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   167
val spec'= read_instantiate [("x","P::?'b=>bool")] spec;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   168
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   169
fun tracing true _ = ()
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   170
  | tracing false msg = writeln msg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   171
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   172
fun simplify_defn strict thy cs ss congs wfs id pats def0 =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   173
   let val def = Thm.freezeT def0 RS meta_eq_to_obj_eq
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   174
       val {rules,rows,TCs,full_pats_TCs} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   175
           Prim.post_definition congs (thy, (def,pats))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   176
       val {lhs=f,rhs} = S.dest_eq (concl def)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   177
       val (_,[R,_]) = S.strip_comb rhs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   178
       val dummy = Prim.trace_thms "congs =" congs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   179
       (*the next step has caused simplifier looping in some cases*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   180
       val {induction, rules, tcs} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   181
             proof_stage strict cs ss wfs thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   182
               {f = f, R = R, rules = rules,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   183
                full_pats_TCs = full_pats_TCs,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   184
                TCs = TCs}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   185
       val rules' = map (standard o ObjectLogic.rulify_no_asm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   186
                        (R.CONJUNCTS rules)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   187
         in  {induct = meta_outer (ObjectLogic.rulify_no_asm (induction RS spec')),
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   188
        rules = ListPair.zip(rules', rows),
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   189
        tcs = (termination_goals rules') @ tcs}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   190
   end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   191
  handle U.ERR {mesg,func,module} =>
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   192
               error (mesg ^
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   193
                      "\n    (In TFL function " ^ module ^ "." ^ func ^ ")");
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   194
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   195
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   196
(* Derive the initial equations from the case-split rules to meet the
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   197
users specification of the recursive function. 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   198
 Note: We don't do this if the wf conditions fail to be solved, as each
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   199
case may have a different wf condition. We could group the conditions
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   200
together and say that they must be true to solve the general case,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   201
but that would hide from the user which sub-case they were related
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   202
to. Probably this is not important, and it would work fine, but, for now, I
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   203
prefer leaving more fine-grain control to the user. 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   204
-- Lucas Dixon, Aug 2004 *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   205
local
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   206
  fun get_related_thms i = 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   207
      List.mapPartial ((fn (r,x) => if x = i then SOME r else NONE));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   208
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   209
  fun solve_eq (th, [], i) = 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   210
        error "derive_init_eqs: missing rules"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   211
    | solve_eq (th, [a], i) = [(a, i)]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   212
    | solve_eq (th, splitths as (_ :: _), i) = 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   213
      (writeln "Proving unsplit equation...";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   214
      [((standard o ObjectLogic.rulify_no_asm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   215
          (CaseSplit.splitto splitths th), i)])
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   216
      (* if there's an error, pretend nothing happened with this definition 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   217
         We should probably print something out so that the user knows...? *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   218
      handle ERROR s => 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   219
             (warning ("recdef (solve_eq): " ^ s); map (fn x => (x,i)) splitths);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   220
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   221
fun derive_init_eqs sgn rules eqs = 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   222
    let 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   223
      val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o HOLogic.mk_Trueprop) 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   224
                      eqs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   225
      fun countlist l = 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   226
          (rev o snd o (Library.foldl (fn ((i,L), e) => (i + 1,(e,i) :: L)))) ((0,[]), l)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   227
    in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   228
      List.concat (map (fn (e,i) => solve_eq (e, (get_related_thms i rules), i))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   229
                (countlist eqths))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   230
    end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   231
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   232
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   233
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   234
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   235
 * Defining a function with an associated termination relation.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   236
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   237
fun define_i strict thy cs ss congs wfs fid R eqs =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   238
  let val {functional,pats} = Prim.mk_functional thy eqs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   239
      val (thy, def) = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   240
      val {induct, rules, tcs} = 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   241
          simplify_defn strict thy cs ss congs wfs fid pats def
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   242
      val rules' = 
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   243
          if strict then derive_init_eqs thy rules eqs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   244
          else rules
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   245
  in (thy, {rules = rules', induct = induct, tcs = tcs}) end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   246
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   247
fun define strict thy cs ss congs wfs fid R seqs =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   248
  define_i strict thy cs ss congs wfs fid (Sign.read_term thy R) (map (Sign.read_term thy) seqs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   249
    handle U.ERR {mesg,...} => error mesg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   250
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   251
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   252
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   253
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   254
 *     Definitions with synthesized termination relation
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   255
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   256
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   257
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   258
fun func_of_cond_eqn tm =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   259
  #1 (S.strip_comb (#lhs (S.dest_eq (#2 (S.strip_forall (#2 (S.strip_imp tm)))))));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   260
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   261
fun defer_i thy congs fid eqs =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   262
 let val {rules,R,theory,full_pats_TCs,SV,...} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   263
             Prim.lazyR_def thy (Sign.base_name fid) congs eqs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   264
     val f = func_of_cond_eqn (concl (R.CONJUNCT1 rules handle U.ERR _ => rules));
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   265
     val dummy = message "Proving induction theorem ...";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   266
     val induction = Prim.mk_induction theory
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   267
                        {fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   268
 in (theory,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   269
     (*return the conjoined induction rule and recursion equations,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   270
       with assumptions remaining to discharge*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   271
     standard (induction RS (rules RS conjI)))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   272
 end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   273
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   274
fun defer thy congs fid seqs =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   275
  defer_i thy congs fid (map (Sign.read_term thy) seqs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   276
    handle U.ERR {mesg,...} => error mesg;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   277
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   278
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   279
end;