TFL/post.sml
author wenzelm
Wed, 25 Oct 2000 18:33:40 +0200
changeset 10334 e5e6070fcef5
parent 10015 8c16ec5ba62b
permissions -rw-r--r--
add \<le> to list of "good" symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
     1
(*  Title:      TFL/post.sml
3302
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     2
    ID:         $Id$
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     3
    Author:     Konrad Slind, Cambridge University Computer Laboratory
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     4
    Copyright   1997  University of Cambridge
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     5
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
     6
Second part of main module (postprocessing of TFL definitions).
3302
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     7
*)
404fe31fd8d2 New headers and other minor changes
paulson
parents: 3271
diff changeset
     8
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
     9
signature TFL =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    10
sig
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    11
  val trace: bool ref
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    12
  val quiet_mode: bool ref
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    13
  val message: string -> unit
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    14
  val tgoalw: theory -> thm list -> thm list -> thm list
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    15
  val tgoal: theory -> thm list -> thm list
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    16
  val std_postprocessor: claset -> simpset -> thm list -> theory ->
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    17
    {induction: thm, rules: thm, TCs: term list list} ->
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    18
    {induction: thm, rules: thm, nested_tcs: thm list}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    19
  val define_i: theory -> claset -> simpset -> thm list -> thm list -> xstring ->
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    20
    term -> term list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    21
  val define: theory -> claset -> simpset -> thm list -> thm list -> xstring ->
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    22
    string -> string list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    23
  val defer_i: theory -> thm list -> xstring -> term list -> theory * thm
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    24
  val defer: theory -> thm list -> xstring -> string list -> theory * thm
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    25
end;
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    26
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    27
structure Tfl: TFL =
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    28
struct
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    29
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    30
structure S = USyntax
8817
1c48b3732399 tuned messages;
wenzelm
parents: 8712
diff changeset
    31
1c48b3732399 tuned messages;
wenzelm
parents: 8712
diff changeset
    32
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    33
(* messages *)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    34
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    35
val trace = Prim.trace
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    36
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    37
val quiet_mode = ref false;
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    38
fun message s = if ! quiet_mode then () else writeln s;
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    39
6524
13410ecfce0b proper quiet_mode;
wenzelm
parents: 6498
diff changeset
    40
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    41
(* misc *)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    42
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    43
fun read_term thy = Sign.simple_read_term (Theory.sign_of thy) HOLogic.termT;
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    44
8817
1c48b3732399 tuned messages;
wenzelm
parents: 8712
diff changeset
    45
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    46
(*---------------------------------------------------------------------------
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    47
 * Extract termination goals so that they can be put it into a goalstack, or
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    48
 * have a tactic directly applied to them.
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    49
 *--------------------------------------------------------------------------*)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    50
fun termination_goals rules =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    51
    map (#1 o Type.freeze_thaw o HOLogic.dest_Trueprop)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    52
      (foldr (fn (th,A) => union_term (prems_of th, A)) (rules, []));
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    53
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    54
(*---------------------------------------------------------------------------
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    55
 * Finds the termination conditions in (highly massaged) definition and
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    56
 * puts them into a goalstack.
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    57
 *--------------------------------------------------------------------------*)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    58
fun tgoalw thy defs rules =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    59
  case termination_goals rules of
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    60
      [] => error "tgoalw: no termination conditions to prove"
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    61
    | L  => goalw_cterm defs
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    62
              (Thm.cterm_of (Theory.sign_of thy)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    63
                        (HOLogic.mk_Trueprop(USyntax.list_mk_conj L)));
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    64
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    65
fun tgoal thy = tgoalw thy [];
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    66
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    67
(*---------------------------------------------------------------------------
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    68
 * Three postprocessors are applied to the definition.  It
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    69
 * attempts to prove wellfoundedness of the given relation, simplifies the
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    70
 * non-proved termination conditions, and finally attempts to prove the
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
    71
 * simplified termination conditions.
3405
2cccd0e3e9ea Numerous simplifications and removal of HOL-isms
paulson
parents: 3391
diff changeset
    72
 *--------------------------------------------------------------------------*)
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    73
fun std_postprocessor cs ss wfs =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    74
  Prim.postprocess
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    75
   {wf_tac     = REPEAT (ares_tac wfs 1),
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    76
    terminator = asm_simp_tac ss 1
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    77
                 THEN TRY (fast_tac (cs addSDs [not0_implies_Suc] addss ss) 1),
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    78
    simplifier = Rules.simpl_conv ss []};
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    79
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    80
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    81
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    82
val concl = #2 o Rules.dest_thm;
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    83
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    84
(*---------------------------------------------------------------------------
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    85
 * Postprocess a definition made by "define". This is a separate stage of
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
    86
 * processing from the definition stage.
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    87
 *---------------------------------------------------------------------------*)
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    88
local
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    89
structure R = Rules
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    90
structure U = Utils
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    91
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    92
(* The rest of these local definitions are for the tricky nested case *)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    93
val solved = not o U.can S.dest_eq o #2 o S.strip_forall o concl
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    94
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    95
fun id_thm th =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    96
   let val {lhs,rhs} = S.dest_eq(#2(S.strip_forall(#2 (R.dest_thm th))))
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
    97
   in lhs aconv rhs
10015
8c16ec5ba62b indicate occurrences of 'handle _';
wenzelm
parents: 9904
diff changeset
    98
   end handle _ => false (* FIXME do not handle _ !!! *)
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
    99
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   100
fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   101
val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   102
val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   103
fun mk_meta_eq r = case concl_of r of
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   104
     Const("==",_)$_$_ => r
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   105
  |   _ $(Const("op =",_)$_$_) => r RS eq_reflection
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   106
  |   _ => r RS P_imp_P_eq_True
3405
2cccd0e3e9ea Numerous simplifications and removal of HOL-isms
paulson
parents: 3391
diff changeset
   107
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   108
(*Is this the best way to invoke the simplifier??*)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   109
fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L))
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   110
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   111
fun join_assums th =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   112
  let val {sign,...} = rep_thm th
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   113
      val tych = cterm_of sign
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   114
      val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th)))
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   115
      val cntxtl = (#1 o S.strip_imp) lhs  (* cntxtl should = cntxtr *)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   116
      val cntxtr = (#1 o S.strip_imp) rhs  (* but union is solider *)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   117
      val cntxt = gen_union (op aconv) (cntxtl, cntxtr)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   118
  in
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   119
    R.GEN_ALL
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   120
      (R.DISCH_ALL
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   121
         (rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th)))
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   122
  end
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   123
  val gen_all = S.gen_all
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   124
in
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   125
fun proof_stage cs ss wfs theory {f, R, rules, full_pats_TCs, TCs} =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   126
  let
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   127
    val _ = message "Proving induction theorem ..."
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   128
    val ind = Prim.mk_induction theory {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   129
    val _ = message "Postprocessing ...";
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   130
    val {rules, induction, nested_tcs} =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   131
      std_postprocessor cs ss wfs theory {rules=rules, induction=ind, TCs=TCs}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   132
  in
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   133
  case nested_tcs
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   134
  of [] => {induction=induction, rules=rules,tcs=[]}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   135
  | L  => let val dummy = message "Simplifying nested TCs ..."
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   136
              val (solved,simplified,stubborn) =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   137
               U.itlist (fn th => fn (So,Si,St) =>
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   138
                     if (id_thm th) then (So, Si, th::St) else
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   139
                     if (solved th) then (th::So, Si, St)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   140
                     else (So, th::Si, St)) nested_tcs ([],[],[])
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   141
              val simplified' = map join_assums simplified
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   142
              val rewr = full_simplify (ss addsimps (solved @ simplified'));
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   143
              val induction' = rewr induction
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   144
              and rules'     = rewr rules
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   145
          in
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   146
          {induction = induction',
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   147
               rules = rules',
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   148
                 tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   149
                           (simplified@stubborn)}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   150
          end
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   151
  end;
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   152
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   153
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   154
(*lcp: curry the predicate of the induction rule*)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   155
fun curry_rule rl = split_rule_var
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   156
                        (head_of (HOLogic.dest_Trueprop (concl_of rl)),
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   157
                         rl);
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   158
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   159
(*lcp: put a theorem into Isabelle form, using meta-level connectives*)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   160
val meta_outer =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   161
    curry_rule o standard o
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   162
    rule_by_tactic (REPEAT
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   163
                    (FIRSTGOAL (resolve_tac [allI, impI, conjI]
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   164
                                ORELSE' etac conjE)));
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   165
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   166
(*Strip off the outer !P*)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   167
val spec'= read_instantiate [("x","P::?'b=>bool")] spec;
3459
112cbb8301dc Removal of structure Context and its replacement by a theorem list of
paulson
parents: 3405
diff changeset
   168
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   169
fun simplify_defn thy cs ss congs wfs id pats def0 =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   170
   let val def = freezeT def0 RS meta_eq_to_obj_eq
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   171
       val {theory,rules,rows,TCs,full_pats_TCs} = Prim.post_definition congs (thy, (def,pats))
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   172
       val {lhs=f,rhs} = S.dest_eq (concl def)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   173
       val (_,[R,_]) = S.strip_comb rhs
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   174
       val {induction, rules, tcs} =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   175
             proof_stage cs ss wfs theory
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   176
               {f = f, R = R, rules = rules,
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   177
                full_pats_TCs = full_pats_TCs,
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   178
                TCs = TCs}
9904
09253f667beb use Rulify.rulify_no_asm;
wenzelm
parents: 9866
diff changeset
   179
       val rules' = map (standard o Rulify.rulify_no_asm) (R.CONJUNCTS rules)
09253f667beb use Rulify.rulify_no_asm;
wenzelm
parents: 9866
diff changeset
   180
   in  {induct = meta_outer (Rulify.rulify_no_asm (induction RS spec')),
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   181
        rules = ListPair.zip(rules', rows),
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   182
        tcs = (termination_goals rules') @ tcs}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   183
   end
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   184
  handle Utils.ERR {mesg,func,module} =>
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   185
               error (mesg ^
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   186
                      "\n    (In TFL function " ^ module ^ "." ^ func ^ ")");
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   187
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   188
(*---------------------------------------------------------------------------
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   189
 * Defining a function with an associated termination relation.
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   190
 *---------------------------------------------------------------------------*)
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   191
fun define_i thy cs ss congs wfs fid R eqs =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   192
  let val {functional,pats} = Prim.mk_functional thy eqs
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   193
      val (thy, def) = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   194
  in (thy, simplify_defn thy cs ss congs wfs fid pats def) end;
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   195
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   196
fun define thy cs ss congs wfs fid R seqs =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   197
  define_i thy cs ss congs wfs fid (read_term thy R) (map (read_term thy) seqs)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   198
    handle Utils.ERR {mesg,...} => error mesg;
7262
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   199
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   200
a05dc63ca29b from Konrad: support for schematic definitions
paulson
parents: 6554
diff changeset
   201
(*---------------------------------------------------------------------------
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   202
 *
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   203
 *     Definitions with synthesized termination relation
3191
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   204
 *
14bd6e5985f1 TFL now integrated with HOL (more work needed)
paulson
parents: 2467
diff changeset
   205
 *---------------------------------------------------------------------------*)
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   206
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   207
local open USyntax
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   208
in
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   209
fun func_of_cond_eqn tm =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   210
  #1(strip_comb(#lhs(dest_eq(#2 (strip_forall(#2(strip_imp tm)))))))
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   211
end;
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   212
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   213
fun defer_i thy congs fid eqs =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   214
 let val {rules,R,theory,full_pats_TCs,SV,...} =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   215
             Prim.lazyR_def thy (Sign.base_name fid) congs eqs
10015
8c16ec5ba62b indicate occurrences of 'handle _';
wenzelm
parents: 9904
diff changeset
   216
     val f = func_of_cond_eqn (concl(R.CONJUNCT1 rules handle _ => rules)) (* FIXME do not handle _ !!! *)
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   217
     val dummy = message "Proving induction theorem ...";
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   218
     val induction = Prim.mk_induction theory
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   219
                        {fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs}
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   220
 in (theory,
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   221
     (*return the conjoined induction rule and recursion equations,
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   222
       with assumptions remaining to discharge*)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   223
     standard (induction RS (rules RS conjI)))
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   224
 end
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   225
9866
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   226
fun defer thy congs fid seqs =
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   227
  defer_i thy congs fid (map (read_term thy) seqs)
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   228
    handle Utils.ERR {mesg,...} => error mesg;
90cbf68b9227 proper handling of hints;
wenzelm
parents: 9736
diff changeset
   229
end;
6498
1ebbe18fe236 Now for recdefs that omit the WF relation;
paulson
parents: 6476
diff changeset
   230
2112
3902e9af752f Konrad Slind's TFL
paulson
parents:
diff changeset
   231
end;