src/HOL/Tools/TFL/tfl.ML
author wenzelm
Mon, 01 Jun 2015 17:08:47 +0200
changeset 60334 63f25a1adcfc
parent 60329 e85ed7a36b2f
child 60335 edac62cd7bde
permissions -rw-r--r--
clarified context;
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/tfl.ML
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Konrad Slind, Cambridge University Computer Laboratory
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     3
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     4
First part of main module.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     5
*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     6
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     7
signature PRIM =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
     8
sig
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
     9
  val trace: bool Unsynchronized.ref
55236
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
    10
  val trace_thms: Proof.context -> string -> thm list -> unit
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
    11
  val trace_cterm: Proof.context -> string -> cterm -> unit
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    12
  type pattern
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    13
  val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    14
  val wfrec_definition0: theory -> string -> term -> term -> theory * thm
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
    15
  val post_definition: thm list -> theory -> Proof.context -> thm * pattern list ->
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    16
   {rules: thm,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    17
    rows: int list,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    18
    TCs: term list list,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    19
    full_pats_TCs: (term * term list) list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    20
  val wfrec_eqns: theory -> xstring -> thm list -> term list ->
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    21
   {WFR: term,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    22
    SV: term list,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    23
    proto_def: term,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    24
    extracta: (thm * term list) list,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    25
    pats: pattern list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    26
  val lazyR_def: theory -> xstring -> thm list -> term list ->
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    27
   {theory: theory,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    28
    rules: thm,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    29
    R: term,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    30
    SV: term list,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    31
    full_pats_TCs: (term * term list) list,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    32
    patterns : pattern list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    33
  val mk_induction: theory ->
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    34
    {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    35
  val postprocess: bool -> {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    36
    -> theory -> {rules: thm, induction: thm, TCs: term list list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    37
    -> {rules: thm, induction: thm, nested_tcs: thm list}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    38
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    39
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    40
structure Prim: PRIM =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    41
struct
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    42
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
    43
val trace = Unsynchronized.ref false;
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    44
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    45
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
    46
fun TFL_ERR func mesg = Utils.ERR {module = "Tfl", func = func, mesg = mesg};
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    47
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
    48
val concl = #2 o Rules.dest_thm;
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
    49
val hyp = #1 o Rules.dest_thm;
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    50
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
    51
val list_mk_type = Utils.end_itlist (curry (op -->));
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    52
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    53
fun front_last [] = raise TFL_ERR "front_last" "empty list"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    54
  | front_last [x] = ([],x)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    55
  | front_last (h::t) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    56
     let val (pref,x) = front_last t
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    57
     in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    58
        (h::pref,x)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    59
     end;
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
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    63
 * The next function is common to pattern-match translation and
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    64
 * proof of completeness of cases for the induction theorem.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    65
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    66
 * The curried function "gvvariant" returns a function to generate distinct
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    67
 * variables that are guaranteed not to be in names.  The names of
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    68
 * the variables go u, v, ..., z, aa, ..., az, ...  The returned
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    69
 * function contains embedded refs!
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    70
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    71
fun gvvariant names =
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
    72
  let val slist = Unsynchronized.ref names
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32432
diff changeset
    73
      val vname = Unsynchronized.ref "u"
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    74
      fun new() =
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36616
diff changeset
    75
         if member (op =) (!slist) (!vname)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    76
         then (vname := Symbol.bump_string (!vname);  new())
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    77
         else (slist := !vname :: !slist;  !vname)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    78
  in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    79
  fn ty => Free(new(), ty)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    80
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    81
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    82
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    83
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    84
 * Used in induction theorem production. This is the simple case of
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    85
 * partitioning up pattern rows by the leading constructor.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    86
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    87
fun ipartition gv (constructors,rows) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    88
  let fun pfail s = raise TFL_ERR "partition.part" s
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    89
      fun part {constrs = [],   rows = [],   A} = rev A
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    90
        | part {constrs = [],   rows = _::_, A} = pfail"extra cases in defn"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    91
        | part {constrs = _::_, rows = [],   A} = pfail"cases missing in defn"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    92
        | part {constrs = c::crst, rows,     A} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    93
          let val (c, T) = dest_Const c
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    94
              val L = binder_types T
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    95
              val (in_group, not_in_group) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    96
               fold_rev (fn (row as (p::rst, rhs)) =>
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    97
                         fn (in_group,not_in_group) =>
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
    98
                  let val (pc,args) = USyntax.strip_comb p
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
    99
                  in if (#1(dest_Const pc) = c)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   100
                     then ((args@rst, rhs)::in_group, not_in_group)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   101
                     else (in_group, row::not_in_group)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   102
                  end)      rows ([],[])
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   103
              val col_types = Utils.take type_of (length L, #1(hd in_group))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   104
          in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   105
          part{constrs = crst, rows = not_in_group,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   106
               A = {constructor = c,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   107
                    new_formals = map gv col_types,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   108
                    group = in_group}::A}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   109
          end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   110
  in part{constrs = constructors, rows = rows, A = []}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   111
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   112
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   113
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   114
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   115
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   116
 * Each pattern carries with it a tag (i,b) where
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   117
 * i is the clause it came from and
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   118
 * b=true indicates that clause was given by the user
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   119
 * (or is an instantiation of a user supplied pattern)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   120
 * b=false --> i = ~1
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   121
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   122
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   123
type pattern = term * (int * bool)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   124
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   125
fun pattern_map f (tm,x) = (f tm, x);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   126
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   127
fun pattern_subst theta = pattern_map (subst_free theta);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   128
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   129
val pat_of = fst;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   130
fun row_of_pat x = fst (snd x);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   131
fun given x = snd (snd x);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   132
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   133
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   134
 * Produce an instance of a constructor, plus genvars for its arguments.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   135
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   136
fun fresh_constr ty_match colty gv c =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   137
  let val (_,Ty) = dest_Const c
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   138
      val L = binder_types Ty
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   139
      and ty = body_type Ty
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   140
      val ty_theta = ty_match ty colty
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   141
      val c' = USyntax.inst ty_theta c
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   142
      val gvars = map (USyntax.inst ty_theta o gv) L
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   143
  in (c', gvars)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   144
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   145
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   146
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   147
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   148
 * Goes through a list of rows and picks out the ones beginning with a
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   149
 * pattern with constructor = name.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   150
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   151
fun mk_group name rows =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   152
  fold_rev (fn (row as ((prfx, p::rst), rhs)) =>
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   153
            fn (in_group,not_in_group) =>
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   154
               let val (pc,args) = USyntax.strip_comb p
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   155
               in if ((#1 (Term.dest_Const pc) = name) handle TERM _ => false)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   156
                  then (((prfx,args@rst), rhs)::in_group, not_in_group)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   157
                  else (in_group, row::not_in_group) end)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   158
      rows ([],[]);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   159
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   160
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   161
 * Partition the rows. Not efficient: we should use hashing.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   162
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   163
fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   164
  | partition gv ty_match
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   165
              (constructors, colty, res_ty, rows as (((prfx,_),_)::_)) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   166
let val fresh = fresh_constr ty_match colty gv
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   167
     fun part {constrs = [],      rows, A} = rev A
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   168
       | part {constrs = c::crst, rows, A} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   169
         let val (c',gvars) = fresh c
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   170
             val (in_group, not_in_group) = mk_group (#1 (dest_Const c')) rows
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   171
             val in_group' =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   172
                 if (null in_group)  (* Constructor not given *)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   173
                 then [((prfx, #2(fresh c)), (USyntax.ARB res_ty, (~1,false)))]
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   174
                 else in_group
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   175
         in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   176
         part{constrs = crst,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   177
              rows = not_in_group,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   178
              A = {constructor = c',
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   179
                   new_formals = gvars,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   180
                   group = in_group'}::A}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   181
         end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   182
in part{constrs=constructors, rows=rows, A=[]}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   183
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   184
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   185
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   186
 * Misc. routines used in mk_case
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   187
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   188
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   189
fun mk_pat (c,l) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   190
  let val L = length (binder_types (type_of c))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   191
      fun build (prfx,tag,plist) =
33955
fff6f11b1f09 curried take/drop
haftmann
parents: 33339
diff changeset
   192
          let val (args, plist') = chop L plist
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   193
          in (prfx,tag,list_comb(c,args)::plist') end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   194
  in map build l end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   195
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   196
fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   197
  | v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   198
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   199
fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   200
  | v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   201
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   202
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   203
(*----------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   204
 * Translation of pattern terms into nested case expressions.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   205
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   206
 * This performs the translation and also builds the full set of patterns.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   207
 * Thus it supports the construction of induction theorems even when an
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   208
 * incomplete set of patterns is given.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   209
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   210
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   211
fun mk_case ty_info ty_match usednames range_ty =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   212
 let
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   213
 fun mk_case_fail s = raise TFL_ERR "mk_case" s
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   214
 val fresh_var = gvvariant usednames
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   215
 val divide = partition fresh_var ty_match
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   216
 fun expand constructors ty ((_,[]), _) = mk_case_fail"expand_var_row"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   217
   | expand constructors ty (row as ((prfx, p::rst), rhs)) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   218
       if (is_Free p)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   219
       then let val fresh = fresh_constr ty_match ty fresh_var
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   220
                fun expnd (c,gvs) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   221
                  let val capp = list_comb(c,gvs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   222
                  in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   223
                  end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   224
            in map expnd (map fresh constructors)  end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   225
       else [row]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   226
 fun mk{rows=[],...} = mk_case_fail"no rows"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   227
   | mk{path=[], rows = ((prfx, []), (tm,tag))::_} =  (* Done *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   228
        ([(prfx,tag,[])], tm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   229
   | mk{path=[], rows = _::_} = mk_case_fail"blunder"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   230
   | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   231
        mk{path = path,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   232
           rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   233
   | mk{path = u::rstp, rows as ((_, p::_), _)::_} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   234
     let val (pat_rectangle,rights) = ListPair.unzip rows
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   235
         val col0 = map(hd o #2) pat_rectangle
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   236
     in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   237
     if (forall is_Free col0)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   238
     then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   239
                                (ListPair.zip (col0, rights))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   240
              val pat_rectangle' = map v_to_prfx pat_rectangle
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   241
              val (pref_patl,tm) = mk{path = rstp,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   242
                                      rows = ListPair.zip (pat_rectangle',
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   243
                                                           rights')}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   244
          in (map v_to_pats pref_patl, tm)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   245
          end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   246
     else
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   247
     let val pty as Type (ty_name,_) = type_of p
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   248
     in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   249
     case (ty_info ty_name)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   250
     of NONE => mk_case_fail("Not a known datatype: "^ty_name)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   251
      | SOME{case_const,constructors} =>
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   252
        let
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   253
            val case_const_name = #1(dest_Const case_const)
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   254
            val nrows = maps (expand constructors pty) rows
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   255
            val subproblems = divide(constructors, pty, range_ty, nrows)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   256
            val groups      = map #group subproblems
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   257
            and new_formals = map #new_formals subproblems
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   258
            and constructors' = map #constructor subproblems
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   259
            val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows})
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   260
                           (ListPair.zip (new_formals, groups))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   261
            val rec_calls = map mk news
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   262
            val (pat_rect,dtrees) = ListPair.unzip rec_calls
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   263
            val case_functions = map USyntax.list_mk_abs
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   264
                                  (ListPair.zip (new_formals, dtrees))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   265
            val types = map type_of (case_functions@[u]) @ [range_ty]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   266
            val case_const' = Const(case_const_name, list_mk_type types)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   267
            val tree = list_comb(case_const', case_functions@[u])
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   268
            val pat_rect1 = flat (ListPair.map mk_pat (constructors', pat_rect))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   269
        in (pat_rect1,tree)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   270
        end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   271
     end end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   272
 in mk
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   273
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   274
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   275
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   276
(* Repeated variable occurrences in a pattern are not allowed. *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   277
fun FV_multiset tm =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   278
   case (USyntax.dest_term tm)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   279
     of USyntax.VAR{Name = c, Ty = T} => [Free(c, T)]
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   280
      | USyntax.CONST _ => []
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   281
      | USyntax.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   282
      | USyntax.LAMB _ => raise TFL_ERR "FV_multiset" "lambda";
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   283
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   284
fun no_repeat_vars thy pat =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   285
 let fun check [] = true
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   286
       | check (v::rst) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   287
         if member (op aconv) rst v then
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   288
            raise TFL_ERR "no_repeat_vars"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   289
                          (quote (#1 (dest_Free v)) ^
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   290
                          " occurs repeatedly in the pattern " ^
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   291
                          quote (Syntax.string_of_term_global thy pat))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   292
         else check rst
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   293
 in check (FV_multiset pat)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   294
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   295
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   296
fun dest_atom (Free p) = p
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   297
  | dest_atom (Const p) = p
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   298
  | dest_atom  _ = raise TFL_ERR "dest_atom" "function name not an identifier";
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   299
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   300
fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   301
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   302
local fun mk_functional_err s = raise TFL_ERR "mk_functional" s
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   303
      fun single [_$_] =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   304
              mk_functional_err "recdef does not allow currying"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   305
        | single [f] = f
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   306
        | single fs  =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   307
              (*multiple function names?*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   308
              if length (distinct same_name fs) < length fs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   309
              then mk_functional_err
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   310
                   "The function being declared appears with multiple types"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   311
              else mk_functional_err
41491
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41164
diff changeset
   312
                   (string_of_int (length fs) ^
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   313
                    " distinct function names being declared")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   314
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   315
fun mk_functional thy clauses =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   316
 let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   317
                   handle TERM _ => raise TFL_ERR "mk_functional"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   318
                        "recursion equations must use the = relation")
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   319
     val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   320
     val atom = single (distinct (op aconv) funcs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   321
     val (fname,ftype) = dest_atom atom
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   322
     val dummy = map (no_repeat_vars thy) pats
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   323
     val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats,
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 33043
diff changeset
   324
                              map_index (fn (i, t) => (t,(i,true))) R)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43324
diff changeset
   325
     val names = List.foldr Misc_Legacy.add_term_names [] R
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   326
     val atype = type_of(hd pats)
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   327
     and aname = singleton (Name.variant_list names) "a"
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   328
     val a = Free(aname,atype)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   329
     val ty_info = Thry.match_info thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   330
     val ty_match = Thry.match_type thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   331
     val range_ty = type_of (hd R)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   332
     val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   333
                                    {path=[a], rows=rows}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   334
     val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   335
          handle Match => mk_functional_err "error in pattern-match translation"
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58354
diff changeset
   336
     val patts2 = Library.sort (Library.int_ord o apply2 row_of_pat) patts1
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   337
     val finals = map row_of_pat patts2
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   338
     val originals = map (row_of_pat o #2) rows
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 33039
diff changeset
   339
     val dummy = case (subtract (op =) finals originals)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   340
             of [] => ()
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   341
          | L => mk_functional_err
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   342
 ("The following clauses are redundant (covered by preceding clauses): " ^
41491
a2ad5b824051 eliminated Int.toString;
wenzelm
parents: 41164
diff changeset
   343
                   commas (map (fn i => string_of_int (i + 1)) L))
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
   344
 in {functional = Abs(Long_Name.base_name fname, ftype,
44241
7943b69f0188 modernized signature of Term.absfree/absdummy;
wenzelm
parents: 44121
diff changeset
   345
                      abstract_over (atom, absfree (aname,atype) case_tm)),
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   346
     pats = patts2}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   347
end end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   348
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   349
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   350
(*----------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   351
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   352
 *                    PRINCIPLES OF DEFINITION
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   353
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   354
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   355
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   356
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   357
(*For Isabelle, the lhs of a definition must be a constant.*)
35799
7adb03f27b28 preserve full const name more carefully, and avoid slightly odd Sign.intern_term;
wenzelm
parents: 35666
diff changeset
   358
fun const_def sign (c, Ty, rhs) =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41491
diff changeset
   359
  singleton (Syntax.check_terms (Proof_Context.init_global sign))
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55236
diff changeset
   360
    (Const(@{const_name Pure.eq},dummyT) $ Const(c,Ty) $ rhs);
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   361
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   362
(*Make all TVars available for instantiation by adding a ? to the front*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   363
fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   364
  | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   365
  | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort);
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   366
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   367
local val f_eq_wfrec_R_M =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   368
           #ant(USyntax.dest_imp(#2(USyntax.strip_forall (concl Thms.WFREC_COROLLARY))))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   369
      val {lhs=f, rhs} = USyntax.dest_eq f_eq_wfrec_R_M
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   370
      val (fname,_) = dest_Free f
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   371
      val (wfrec,_) = USyntax.strip_comb rhs
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   372
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   373
fun wfrec_definition0 thy fid R (functional as Abs(x, Ty, _)) =
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 45620
diff changeset
   374
 let val def_name = Thm.def_name (Long_Name.base_name fid)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   375
     val wfrec_R_M =  map_types poly_tvars
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   376
                          (wfrec $ map_types poly_tvars R)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   377
                      $ functional
35799
7adb03f27b28 preserve full const name more carefully, and avoid slightly odd Sign.intern_term;
wenzelm
parents: 35666
diff changeset
   378
     val def_term = const_def thy (fid, Ty, wfrec_R_M)
7adb03f27b28 preserve full const name more carefully, and avoid slightly odd Sign.intern_term;
wenzelm
parents: 35666
diff changeset
   379
     val ([def], thy') =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39158
diff changeset
   380
      Global_Theory.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   381
 in (thy', def) end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   382
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   383
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   384
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   385
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   386
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   387
 * This structure keeps track of congruence rules that aren't derived
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   388
 * from a datatype definition.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   389
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   390
fun extraction_thms thy =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   391
 let val {case_rewrites,case_congs} = Thry.extract_info thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   392
 in (case_rewrites, case_congs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   393
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   394
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   395
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   396
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   397
 * Pair patterns with termination conditions. The full list of patterns for
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   398
 * a definition is merged with the TCs arising from the user-given clauses.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   399
 * There can be fewer clauses than the full list, if the user omitted some
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   400
 * cases. This routine is used to prepare input for mk_induction.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   401
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   402
fun merge full_pats TCs =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   403
let fun insert (p,TCs) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   404
      let fun insrt ((x as (h,[]))::rst) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   405
                 if (p aconv h) then (p,TCs)::rst else x::insrt rst
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   406
            | insrt (x::rst) = x::insrt rst
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   407
            | insrt[] = raise TFL_ERR "merge.insert" "pattern not found"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   408
      in insrt end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   409
    fun pass ([],ptcl_final) = ptcl_final
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   410
      | pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   411
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   412
  pass (TCs, map (fn p => (p,[])) full_pats)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   413
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   414
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   415
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33063
diff changeset
   416
fun givens pats = map pat_of (filter given pats);
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   417
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   418
fun post_definition meta_tflCongs theory ctxt (def, pats) =
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   419
 let val tych = Thry.typecheck theory
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   420
     val f = #lhs(USyntax.dest_eq(concl def))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   421
     val corollary = Rules.MATCH_MP Thms.WFREC_COROLLARY def
33317
b4534348b8fd standardized filter/filter_out;
wenzelm
parents: 33063
diff changeset
   422
     val pats' = filter given pats
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   423
     val given_pats = map pat_of pats'
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   424
     val rows = map row_of_pat pats'
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   425
     val WFR = #ant(USyntax.dest_imp(concl corollary))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   426
     val R = #Rand(USyntax.dest_comb WFR)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   427
     val corollary' = Rules.UNDISCH corollary  (* put WF R on assums *)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   428
     val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary')
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   429
                           given_pats
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   430
     val (case_rewrites,context_congs) = extraction_thms theory
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   431
     (*case_ss causes minimal simplification: bodies of case expressions are
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   432
       not simplified. Otherwise large examples (Red-Black trees) are too
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   433
       slow.*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46909
diff changeset
   434
     val case_simpset =
54999
7859ed58c041 clarified context;
wenzelm
parents: 54895
diff changeset
   435
       put_simpset HOL_basic_ss ctxt
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46909
diff changeset
   436
          addsimps case_rewrites
57983
6edc3529bb4e reordered some (co)datatype property names for more consistency
blanchet
parents: 56245
diff changeset
   437
          |> fold (Simplifier.add_cong o #case_cong_weak o snd)
58354
04ac60da613e support (finite values of) codatatypes in Quickcheck
blanchet
parents: 58132
diff changeset
   438
              (Symtab.dest (BNF_LFP_Compat.get_all theory [BNF_LFP_Compat.Keep_Nesting]))
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46909
diff changeset
   439
     val corollaries' = map (Simplifier.simplify case_simpset) corollaries
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   440
     val extract = Rules.CONTEXT_REWRITE_RULE
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 46909
diff changeset
   441
                     (f, [R], @{thm cut_apply}, meta_tflCongs @ context_congs)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   442
     val (rules, TCs) = ListPair.unzip (map extract corollaries')
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   443
     val rules0 = map (rewrite_rule ctxt [Thms.CUT_DEF]) rules
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   444
     val mk_cond_rule = Rules.FILTER_DISCH_ALL(not o curry (op aconv) WFR)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   445
     val rules1 = Rules.LIST_CONJ(map mk_cond_rule rules0)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   446
 in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   447
 {rules = rules1,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   448
  rows = rows,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   449
  full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)),
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   450
  TCs = TCs}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   451
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   452
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   453
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   454
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   455
 * Perform the extraction without making the definition. Definition and
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   456
 * extraction commute for the non-nested case.  (Deferred recdefs)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   457
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   458
 * The purpose of wfrec_eqns is merely to instantiate the recursion theorem
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   459
 * and extract termination conditions: no definition is made.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   460
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   461
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   462
fun wfrec_eqns thy fid tflCongs eqns =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   463
 let val ctxt = Proof_Context.init_global thy
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   464
     val {lhs,rhs} = USyntax.dest_eq (hd eqns)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   465
     val (f,args) = USyntax.strip_comb lhs
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   466
     val (fname,fty) = dest_atom f
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   467
     val (SV,a) = front_last args    (* SV = schematic variables *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   468
     val g = list_comb(f,SV)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   469
     val h = Free(fname,type_of g)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   470
     val eqns1 = map (subst_free[(g,h)]) eqns
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   471
     val {functional as Abs(x, Ty, _),  pats} = mk_functional thy eqns1
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   472
     val given_pats = givens pats
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   473
     (* val f = Free(x,Ty) *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   474
     val Type("fun", [f_dty, f_rty]) = Ty
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   475
     val dummy = if x<>fid then
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   476
                        raise TFL_ERR "wfrec_eqns"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   477
                                      ("Expected a definition of " ^
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   478
                                      quote fid ^ " but found one of " ^
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   479
                                      quote x)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   480
                 else ()
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   481
     val (case_rewrites,context_congs) = extraction_thms thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   482
     val tych = Thry.typecheck thy
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   483
     val WFREC_THM0 = Rules.ISPEC (tych functional) Thms.WFREC_COROLLARY
38554
f8999e19dd49 corrected some long-overseen misperceptions in recdef
haftmann
parents: 38549
diff changeset
   484
     val Const(@{const_name All},_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43324
diff changeset
   485
     val R = Free (singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] eqns)) Rname,
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   486
                   Rtype)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   487
     val WFREC_THM = Rules.ISPECL [tych R, tych g] WFREC_THM0
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   488
     val ([proto_def, WFR],_) = USyntax.strip_imp(concl WFREC_THM)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   489
     val dummy =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   490
           if !trace then
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   491
               writeln ("ORIGINAL PROTO_DEF: " ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26928
diff changeset
   492
                          Syntax.string_of_term_global thy proto_def)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   493
           else ()
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   494
     val R1 = USyntax.rand WFR
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   495
     val corollary' = Rules.UNDISCH (Rules.UNDISCH WFREC_THM)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   496
     val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   497
     val corollaries' = map (rewrite_rule ctxt case_rewrites) corollaries
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   498
     fun extract X = Rules.CONTEXT_REWRITE_RULE
26177
6b127c056e83 removed legacy ML bindings;
wenzelm
parents: 24493
diff changeset
   499
                       (f, R1::SV, @{thm cut_apply}, tflCongs@context_congs) X
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   500
 in {proto_def = proto_def,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   501
     SV=SV,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   502
     WFR=WFR,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   503
     pats=pats,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   504
     extracta = map extract corollaries'}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   505
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   506
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   507
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   508
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   509
 * Define the constant after extracting the termination conditions. The
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   510
 * wellfounded relation used in the definition is computed by using the
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   511
 * choice operator on the extracted conditions (plus the condition that
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   512
 * such a relation must be wellfounded).
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   513
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   514
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   515
fun lazyR_def thy fid tflCongs eqns =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   516
 let val {proto_def,WFR,pats,extracta,SV} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   517
           wfrec_eqns thy fid tflCongs eqns
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   518
     val R1 = USyntax.rand WFR
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   519
     val f = #lhs(USyntax.dest_eq proto_def)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   520
     val (extractants,TCl) = ListPair.unzip extracta
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   521
     val dummy = if !trace
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31784
diff changeset
   522
                 then writeln (cat_lines ("Extractants =" ::
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31784
diff changeset
   523
                  map (Display.string_of_thm_global thy) extractants))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   524
                 else ()
33339
d41f77196338 eliminated some old folds;
wenzelm
parents: 33317
diff changeset
   525
     val TCs = fold_rev (union (op aconv)) TCl []
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   526
     val full_rqt = WFR::TCs
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   527
     val R' = USyntax.mk_select{Bvar=R1, Body=USyntax.list_mk_conj full_rqt}
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   528
     val R'abs = USyntax.rand R'
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   529
     val proto_def' = subst_free[(R1,R')] proto_def
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   530
     val dummy = if !trace then writeln ("proto_def' = " ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26928
diff changeset
   531
                                         Syntax.string_of_term_global
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   532
                                         thy proto_def')
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   533
                           else ()
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   534
     val {lhs,rhs} = USyntax.dest_eq proto_def'
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   535
     val (c,args) = USyntax.strip_comb lhs
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   536
     val (name,Ty) = dest_atom c
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   537
     val defn = const_def thy (name, Ty, USyntax.list_mk_abs (args,rhs))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   538
     val ([def0], theory) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   539
       thy
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 39158
diff changeset
   540
       |> Global_Theory.add_defs false
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 45620
diff changeset
   541
            [Thm.no_attributes (Binding.name (Thm.def_name fid), defn)]
36616
712724c32ac8 replaced Thm.legacy_freezeT by Thm.unvarify_global -- these facts stem from closed definitions, i.e. there are no term Vars;
wenzelm
parents: 36615
diff changeset
   542
     val def = Thm.unvarify_global def0;
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31784
diff changeset
   543
     val dummy =
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31784
diff changeset
   544
       if !trace then writeln ("DEF = " ^ Display.string_of_thm_global theory def)
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31784
diff changeset
   545
       else ()
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   546
     (* val fconst = #lhs(USyntax.dest_eq(concl def))  *)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   547
     val tych = Thry.typecheck theory
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   548
     val full_rqt_prop = map (Dcterm.mk_prop o tych) full_rqt
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   549
         (*lcp: a lot of object-logic inference to remove*)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   550
     val baz = Rules.DISCH_ALL
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   551
                 (fold_rev Rules.DISCH full_rqt_prop
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   552
                  (Rules.LIST_CONJ extractants))
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31784
diff changeset
   553
     val dum = if !trace then writeln ("baz = " ^ Display.string_of_thm_global theory baz)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   554
                           else ()
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   555
     val f_free = Free (fid, fastype_of f)  (*'cos f is a Const*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   556
     val SV' = map tych SV;
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   557
     val SVrefls = map Thm.reflexive SV'
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   558
     val def0 = (fold (fn x => fn th => Rules.rbeta(Thm.combination th x))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   559
                   SVrefls def)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   560
                RS meta_eq_to_obj_eq
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   561
     val def' = Rules.MP (Rules.SPEC (tych R') (Rules.GEN (tych R1) baz)) def0
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   562
     val body_th = Rules.LIST_CONJ (map Rules.ASSUME full_rqt_prop)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   563
     val SELECT_AX = (*in this way we hope to avoid a STATIC dependence upon
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   564
                       theory Hilbert_Choice*)
39158
e6b96b4cde7e ML_Context.thm;
wenzelm
parents: 38554
diff changeset
   565
         ML_Context.thm "Hilbert_Choice.tfl_some"
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   566
         handle ERROR msg => cat_error msg
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   567
    "defer_recdef requires theory Main or at least Hilbert_Choice as parent"
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   568
     val bar = Rules.MP (Rules.ISPECL[tych R'abs, tych R1] SELECT_AX) body_th
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   569
 in {theory = theory, R=R1, SV=SV,
60329
e85ed7a36b2f eliminated odd C combinator -- Isabelle/ML usually has canonical argument order;
wenzelm
parents: 59058
diff changeset
   570
     rules = fold (fn a => fn b => Rules.MP b a) (Rules.CONJUNCTS bar) def',
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   571
     full_pats_TCs = merge (map pat_of pats) (ListPair.zip (givens pats, TCl)),
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   572
     patterns = pats}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   573
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   574
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   575
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   576
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   577
(*----------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   578
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   579
 *                           INDUCTION THEOREM
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   580
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   581
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   582
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   583
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   584
(*------------------------  Miscellaneous function  --------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   585
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   586
 *           [x_1,...,x_n]     ?v_1...v_n. M[v_1,...,v_n]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   587
 *     -----------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   588
 *     ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]),
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   589
 *                        ...
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   590
 *                        (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] )
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   591
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   592
 * This function is totally ad hoc. Used in the production of the induction
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   593
 * theorem. The nchotomy theorem can have clauses that look like
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   594
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   595
 *     ?v1..vn. z = C vn..v1
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   596
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   597
 * in which the order of quantification is not the order of occurrence of the
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   598
 * quantified variables as arguments to C. Since we have no control over this
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   599
 * aspect of the nchotomy theorem, we make the correspondence explicit by
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   600
 * pairing the incoming new variable with the term it gets beta-reduced into.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   601
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   602
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   603
fun alpha_ex_unroll (xlist, tm) =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   604
  let val (qvars,body) = USyntax.strip_exists tm
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   605
      val vlist = #2 (USyntax.strip_comb (USyntax.rhs body))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   606
      val plist = ListPair.zip (vlist, xlist)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   607
      val args = map (the o AList.lookup (op aconv) plist) qvars
51930
52fd62618631 prefer explicitly qualified exceptions, which is particular important for robust handlers;
wenzelm
parents: 51717
diff changeset
   608
                   handle Option.Option => raise Fail "TFL.alpha_ex_unroll: no correspondence"
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   609
      fun build ex      []   = []
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   610
        | build (_$rex) (v::rst) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   611
           let val ex1 = Term.betapply(rex, v)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   612
           in  ex1 :: build ex1 rst
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   613
           end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   614
     val (nex::exl) = rev (tm::build tm args)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   615
  in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   616
  (nex, ListPair.zip (args, rev exl))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   617
  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   618
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   619
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   620
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   621
(*----------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   622
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   623
 *             PROVING COMPLETENESS OF PATTERNS
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   624
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   625
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   626
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   627
fun mk_case ty_info usednames thy =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   628
 let
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   629
 val ctxt = Proof_Context.init_global thy
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   630
 val divide = ipartition (gvvariant usednames)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   631
 val tych = Thry.typecheck thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   632
 fun tych_binding(x,y) = (tych x, tych y)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   633
 fun fail s = raise TFL_ERR "mk_case" s
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   634
 fun mk{rows=[],...} = fail"no rows"
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   635
   | mk{path=[], rows = [([], (thm, bindings))]} =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   636
                         Rules.IT_EXISTS (map tych_binding bindings) thm
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   637
   | mk{path = u::rstp, rows as (p::_, _)::_} =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   638
     let val (pat_rectangle,rights) = ListPair.unzip rows
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   639
         val col0 = map hd pat_rectangle
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   640
         val pat_rectangle' = map tl pat_rectangle
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   641
     in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   642
     if (forall is_Free col0) (* column 0 is all variables *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   643
     then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)]))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   644
                                (ListPair.zip (rights, col0))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   645
          in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')}
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   646
          end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   647
     else                     (* column 0 is all constructors *)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   648
     let val Type (ty_name,_) = type_of p
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   649
     in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   650
     case (ty_info ty_name)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   651
     of NONE => fail("Not a known datatype: "^ty_name)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   652
      | SOME{constructors,nchotomy} =>
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   653
        let val thm' = Rules.ISPEC (tych u) nchotomy
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   654
            val disjuncts = USyntax.strip_disj (concl thm')
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   655
            val subproblems = divide(constructors, rows)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   656
            val groups      = map #group subproblems
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   657
            and new_formals = map #new_formals subproblems
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   658
            val existentials = ListPair.map alpha_ex_unroll
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   659
                                   (new_formals, disjuncts)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   660
            val constraints = map #1 existentials
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   661
            val vexl = map #2 existentials
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   662
            fun expnd tm (pats,(th,b)) = (pats, (Rules.SUBS ctxt [Rules.ASSUME (tych tm)] th, b))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   663
            val news = map (fn (nf,rows,c) => {path = nf@rstp,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   664
                                               rows = map (expnd c) rows})
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   665
                           (Utils.zip3 new_formals groups constraints)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   666
            val recursive_thms = map mk news
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   667
            val build_exists = Library.foldr
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   668
                                (fn((x,t), th) =>
60334
63f25a1adcfc clarified context;
wenzelm
parents: 60329
diff changeset
   669
                                 Rules.CHOOSE ctxt (tych x, Rules.ASSUME (tych t)) th)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   670
            val thms' = ListPair.map build_exists (vexl, recursive_thms)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   671
            val same_concls = Rules.EVEN_ORS thms'
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   672
        in Rules.DISJ_CASESL thm' same_concls
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   673
        end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   674
     end end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   675
 in mk
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   676
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   677
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   678
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   679
fun complete_cases thy =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   680
 let val ctxt = Proof_Context.init_global thy
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   681
     val tych = Thry.typecheck thy
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   682
     val ty_info = Thry.induct_info thy
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   683
 in fn pats =>
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43324
diff changeset
   684
 let val names = List.foldr Misc_Legacy.add_term_names [] pats
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   685
     val T = type_of (hd pats)
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   686
     val aname = singleton (Name.variant_list names) "a"
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   687
     val vname = singleton (Name.variant_list (aname::names)) "v"
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   688
     val a = Free (aname, T)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   689
     val v = Free (vname, T)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   690
     val a_eq_v = HOLogic.mk_eq(a,v)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   691
     val ex_th0 = Rules.EXISTS (tych (USyntax.mk_exists{Bvar=v,Body=a_eq_v}), tych a)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   692
                           (Rules.REFL (tych a))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   693
     val th0 = Rules.ASSUME (tych a_eq_v)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   694
     val rows = map (fn x => ([x], (th0,[]))) pats
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   695
 in
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   696
 Rules.GEN (tych a)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   697
       (Rules.RIGHT_ASSOC ctxt
60334
63f25a1adcfc clarified context;
wenzelm
parents: 60329
diff changeset
   698
          (Rules.CHOOSE ctxt (tych v, ex_th0)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   699
                (mk_case ty_info (vname::aname::names)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   700
                 thy {path=[v], rows=rows})))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   701
 end end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   702
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   703
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   704
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   705
 * Constructing induction hypotheses: one for each recursive call.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   706
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   707
 * Note. R will never occur as a variable in the ind_clause, because
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   708
 * to do so, it would have to be from a nested definition, and we don't
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   709
 * allow nested defns to have R variable.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   710
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   711
 * Note. When the context is empty, there can be no local variables.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   712
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   713
(*
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   714
local infix 5 ==>
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   715
      fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2}
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   716
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   717
fun build_ih f P (pat,TCs) =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   718
 let val globals = USyntax.free_vars_lr pat
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   719
     fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   720
     fun dest_TC tm =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   721
         let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   722
             val (R,y,_) = USyntax.dest_relation R_y_pat
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   723
             val P_y = if (nested tm) then R_y_pat ==> P$y else P$y
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   724
         in case cntxt
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   725
              of [] => (P_y, (tm,[]))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   726
               | _  => let
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   727
                    val imp = USyntax.list_mk_conj cntxt ==> P_y
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   728
                    val lvs = gen_rems (op aconv) (USyntax.free_vars_lr imp, globals)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   729
                    val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   730
                    in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   731
         end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   732
 in case TCs
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   733
    of [] => (USyntax.list_mk_forall(globals, P$pat), [])
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   734
     |  _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   735
                 val ind_clause = USyntax.list_mk_conj ihs ==> P$pat
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   736
             in (USyntax.list_mk_forall(globals,ind_clause), TCs_locals)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   737
             end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   738
 end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   739
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   740
*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   741
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   742
local infix 5 ==>
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   743
      fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2}
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   744
in
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   745
fun build_ih f (P,SV) (pat,TCs) =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   746
 let val pat_vars = USyntax.free_vars_lr pat
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   747
     val globals = pat_vars@SV
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   748
     fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   749
     fun dest_TC tm =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   750
         let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   751
             val (R,y,_) = USyntax.dest_relation R_y_pat
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   752
             val P_y = if (nested tm) then R_y_pat ==> P$y else P$y
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   753
         in case cntxt
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   754
              of [] => (P_y, (tm,[]))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   755
               | _  => let
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   756
                    val imp = USyntax.list_mk_conj cntxt ==> P_y
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   757
                    val lvs = subtract (op aconv) globals (USyntax.free_vars_lr imp)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   758
                    val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   759
                    in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   760
         end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   761
 in case TCs
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   762
    of [] => (USyntax.list_mk_forall(pat_vars, P$pat), [])
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   763
     |  _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   764
                 val ind_clause = USyntax.list_mk_conj ihs ==> P$pat
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   765
             in (USyntax.list_mk_forall(pat_vars,ind_clause), TCs_locals)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   766
             end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   767
 end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   768
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   769
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   770
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   771
 * This function makes good on the promise made in "build_ih".
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   772
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   773
 * Input  is tm = "(!y. R y pat ==> P y) ==> P pat",
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   774
 *           TCs = TC_1[pat] ... TC_n[pat]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   775
 *           thm = ih1 /\ ... /\ ih_n |- ih[pat]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   776
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   777
fun prove_case f thy (tm,TCs_locals,thm) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   778
 let val tych = Thry.typecheck thy
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   779
     val antc = tych(#ant(USyntax.dest_imp tm))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   780
     val thm' = Rules.SPEC_ALL thm
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   781
     fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   782
     fun get_cntxt TC = tych(#ant(USyntax.dest_imp(#2(USyntax.strip_forall(concl TC)))))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   783
     fun mk_ih ((TC,locals),th2,nested) =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   784
         Rules.GENL (map tych locals)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   785
            (if nested then Rules.DISCH (get_cntxt TC) th2 handle Utils.ERR _ => th2
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   786
             else if USyntax.is_imp (concl TC) then Rules.IMP_TRANS TC th2
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   787
             else Rules.MP th2 TC)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   788
 in
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   789
 Rules.DISCH antc
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   790
 (if USyntax.is_imp(concl thm') (* recursive calls in this clause *)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   791
  then let val th1 = Rules.ASSUME antc
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   792
           val TCs = map #1 TCs_locals
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   793
           val ylist = map (#2 o USyntax.dest_relation o #2 o USyntax.strip_imp o
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   794
                            #2 o USyntax.strip_forall) TCs
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   795
           val TClist = map (fn(TC,lvs) => (Rules.SPEC_ALL(Rules.ASSUME(tych TC)),lvs))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   796
                            TCs_locals
60329
e85ed7a36b2f eliminated odd C combinator -- Isabelle/ML usually has canonical argument order;
wenzelm
parents: 59058
diff changeset
   797
           val th2list = map (fn t => Rules.SPEC (tych t) th1) ylist
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   798
           val nlist = map nested TCs
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   799
           val triples = Utils.zip3 TClist th2list nlist
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   800
           val Pylist = map mk_ih triples
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   801
       in Rules.MP thm' (Rules.LIST_CONJ Pylist) end
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   802
  else thm')
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   803
 end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   804
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   805
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   806
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   807
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   808
 *         x = (v1,...,vn)  |- M[x]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   809
 *    ---------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   810
 *      ?v1 ... vn. x = (v1,...,vn) |- M[x]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   811
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   812
 *---------------------------------------------------------------------------*)
60334
63f25a1adcfc clarified context;
wenzelm
parents: 60329
diff changeset
   813
fun LEFT_ABS_VSTRUCT ctxt tych thm =
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   814
  let fun CHOOSER v (tm,thm) =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   815
        let val ex_tm = USyntax.mk_exists{Bvar=v,Body=tm}
60334
63f25a1adcfc clarified context;
wenzelm
parents: 60329
diff changeset
   816
        in (ex_tm, Rules.CHOOSE ctxt (tych v, Rules.ASSUME (tych ex_tm)) thm)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   817
        end
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   818
      val [veq] = filter (can USyntax.dest_eq) (#1 (Rules.dest_thm thm))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   819
      val {lhs,rhs} = USyntax.dest_eq veq
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   820
      val L = USyntax.free_vars_lr rhs
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   821
  in  #2 (fold_rev CHOOSER L (veq,thm))  end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   822
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   823
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   824
(*----------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   825
 * Input : f, R,  and  [(pat1,TCs1),..., (patn,TCsn)]
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   826
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   827
 * Instantiates WF_INDUCTION_THM, getting Sinduct and then tries to prove
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   828
 * recursion induction (Rinduct) by proving the antecedent of Sinduct from
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   829
 * the antecedent of Rinduct.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   830
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   831
fun mk_induction thy {fconst, R, SV, pat_TCs_list} =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   832
let val ctxt = Proof_Context.init_global thy
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   833
    val tych = Thry.typecheck thy
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   834
    val Sinduction = Rules.UNDISCH (Rules.ISPEC (tych R) Thms.WF_INDUCTION_THM)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   835
    val (pats,TCsl) = ListPair.unzip pat_TCs_list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   836
    val case_thm = complete_cases thy pats
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   837
    val domain = (type_of o hd) pats
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43324
diff changeset
   838
    val Pname = singleton (Name.variant_list (List.foldr (Library.foldr Misc_Legacy.add_term_names)
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   839
                              [] (pats::TCsl))) "P"
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   840
    val P = Free(Pname, domain --> HOLogic.boolT)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   841
    val Sinduct = Rules.SPEC (tych P) Sinduction
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   842
    val Sinduct_assumf = USyntax.rand ((#ant o USyntax.dest_imp o concl) Sinduct)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   843
    val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   844
    val (Rassums,TCl') = ListPair.unzip Rassums_TCl'
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   845
    val Rinduct_assum = Rules.ASSUME (tych (USyntax.list_mk_conj Rassums))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   846
    val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   847
    val tasks = Utils.zip3 cases TCl' (Rules.CONJUNCTS Rinduct_assum)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   848
    val proved_cases = map (prove_case fconst thy) tasks
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   849
    val v =
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   850
      Free (singleton
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 43324
diff changeset
   851
        (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] (map concl proved_cases))) "v",
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42361
diff changeset
   852
          domain)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   853
    val vtyped = tych v
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   854
    val substs = map (Rules.SYM o Rules.ASSUME o tych o (curry HOLogic.mk_eq v)) pats
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51930
diff changeset
   855
    val proved_cases1 = ListPair.map (fn (th,th') => Rules.SUBS ctxt [th]th')
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   856
                          (substs, proved_cases)
60334
63f25a1adcfc clarified context;
wenzelm
parents: 60329
diff changeset
   857
    val abs_cases = map (LEFT_ABS_VSTRUCT ctxt tych) proved_cases1
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   858
    val dant = Rules.GEN vtyped (Rules.DISJ_CASESL (Rules.ISPEC vtyped case_thm) abs_cases)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   859
    val dc = Rules.MP Sinduct dant
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   860
    val Parg_ty = type_of(#Bvar(USyntax.dest_forall(concl dc)))
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   861
    val vars = map (gvvariant[Pname]) (USyntax.strip_prod_type Parg_ty)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   862
    val dc' = fold_rev (Rules.GEN o tych) vars
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   863
                       (Rules.SPEC (tych(USyntax.mk_vstruct Parg_ty vars)) dc)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   864
in
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   865
   Rules.GEN (tych P) (Rules.DISCH (tych(concl Rinduct_assum)) dc')
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   866
end
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   867
handle Utils.ERR _ => raise TFL_ERR "mk_induction" "failed derivation";
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   868
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   869
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   870
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   871
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   872
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   873
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   874
 *                        POST PROCESSING
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   875
 *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   876
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   877
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   878
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   879
fun simplify_induction thy hth ind =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   880
  let val tych = Thry.typecheck thy
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   881
      val (asl,_) = Rules.dest_thm ind
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   882
      val (_,tc_eq_tc') = Rules.dest_thm hth
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   883
      val tc = USyntax.lhs tc_eq_tc'
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   884
      fun loop [] = ind
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   885
        | loop (asm::rst) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   886
          if (can (Thry.match_term thy asm) tc)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   887
          then Rules.UNDISCH
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   888
                 (Rules.MATCH_MP
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   889
                     (Rules.MATCH_MP Thms.simp_thm (Rules.DISCH (tych asm) ind))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   890
                     hth)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   891
         else loop rst
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   892
  in loop asl
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   893
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   894
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   895
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   896
(*---------------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   897
 * The termination condition is an antecedent to the rule, and an
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   898
 * assumption to the theorem.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   899
 *---------------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   900
fun elim_tc tcthm (rule,induction) =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   901
   (Rules.MP rule tcthm, Rules.PROVE_HYP tcthm induction)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   902
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   903
55236
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   904
fun trace_thms ctxt s L =
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   905
  if !trace then writeln (cat_lines (s :: map (Display.string_of_thm ctxt) L))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   906
  else ();
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   907
55236
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   908
fun trace_cterm ctxt s ct =
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   909
  if !trace then
55236
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   910
    writeln (cat_lines [s, Syntax.string_of_term ctxt (Thm.term_of ct)])
32432
64f30bdd3ba1 modernized messages -- eliminated ctyp/cterm operations;
wenzelm
parents: 32091
diff changeset
   911
  else ();
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   912
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   913
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   914
fun postprocess strict {wf_tac, terminator, simplifier} theory {rules,induction,TCs} =
55236
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   915
let val ctxt = Proof_Context.init_global theory;
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   916
    val tych = Thry.typecheck theory;
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   917
    val prove = Rules.prove strict;
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   918
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   919
   (*---------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   920
    * Attempt to eliminate WF condition. It's the only assumption of rules
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   921
    *---------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   922
   val (rules1,induction1)  =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   923
       let val thm = prove(tych(HOLogic.mk_Trueprop
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   924
                                  (hd(#1(Rules.dest_thm rules)))),
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   925
                             wf_tac)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   926
       in (Rules.PROVE_HYP thm rules, Rules.PROVE_HYP thm induction)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   927
       end handle Utils.ERR _ => (rules,induction);
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   928
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   929
   (*----------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   930
    * The termination condition (tc) is simplified to |- tc = tc' (there
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   931
    * might not be a change!) and then 3 attempts are made:
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   932
    *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   933
    *   1. if |- tc = T, then eliminate it with eqT; otherwise,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   934
    *   2. apply the terminator to tc'. If |- tc' = T then eliminate; else
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   935
    *   3. replace tc by tc' in both the rules and the induction theorem.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   936
    *---------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   937
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   938
   fun simplify_tc tc (r,ind) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   939
       let val tc1 = tych tc
55236
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   940
           val _ = trace_cterm ctxt "TC before simplification: " tc1
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   941
           val tc_eq = simplifier tc1
55236
8d61b0aa7a0d proper context for printing;
wenzelm
parents: 54999
diff changeset
   942
           val _ = trace_thms ctxt "result: " [tc_eq]
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   943
       in
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   944
       elim_tc (Rules.MATCH_MP Thms.eqT tc_eq) (r,ind)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   945
       handle Utils.ERR _ =>
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   946
        (elim_tc (Rules.MATCH_MP(Rules.MATCH_MP Thms.rev_eq_mp tc_eq)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   947
                  (prove(tych(HOLogic.mk_Trueprop(USyntax.rhs(concl tc_eq))),
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   948
                           terminator)))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   949
                 (r,ind)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   950
         handle Utils.ERR _ =>
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   951
          (Rules.UNDISCH(Rules.MATCH_MP (Rules.MATCH_MP Thms.simp_thm r) tc_eq),
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   952
           simplify_induction theory tc_eq ind))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   953
       end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   954
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   955
   (*----------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   956
    * Nested termination conditions are harder to get at, since they are
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   957
    * left embedded in the body of the function (and in induction
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   958
    * theorem hypotheses). Our "solution" is to simplify them, and try to
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   959
    * prove termination, but leave the application of the resulting theorem
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   960
    * to a higher level. So things go much as in "simplify_tc": the
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   961
    * termination condition (tc) is simplified to |- tc = tc' (there might
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   962
    * not be a change) and then 2 attempts are made:
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   963
    *
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   964
    *   1. if |- tc = T, then return |- tc; otherwise,
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   965
    *   2. apply the terminator to tc'. If |- tc' = T then return |- tc; else
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   966
    *   3. return |- tc = tc'
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   967
    *---------------------------------------------------------------------*)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   968
   fun simplify_nested_tc tc =
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   969
      let val tc_eq = simplifier (tych (#2 (USyntax.strip_forall tc)))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   970
      in
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   971
      Rules.GEN_ALL
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   972
       (Rules.MATCH_MP Thms.eqT tc_eq
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   973
        handle Utils.ERR _ =>
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   974
          (Rules.MATCH_MP(Rules.MATCH_MP Thms.rev_eq_mp tc_eq)
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   975
                      (prove(tych(HOLogic.mk_Trueprop (USyntax.rhs(concl tc_eq))),
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   976
                               terminator))
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   977
            handle Utils.ERR _ => tc_eq))
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   978
      end
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   979
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   980
   (*-------------------------------------------------------------------
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   981
    * Attempt to simplify the termination conditions in each rule and
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   982
    * in the induction theorem.
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   983
    *-------------------------------------------------------------------*)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   984
   fun strip_imp tm = if USyntax.is_neg tm then ([],tm) else USyntax.strip_imp tm
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   985
   fun loop ([],extras,R,ind) = (rev R, ind, extras)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   986
     | loop ((r,ftcs)::rst, nthms, R, ind) =
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   987
        let val tcs = #1(strip_imp (concl r))
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   988
            val extra_tcs = subtract (op aconv) tcs ftcs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   989
            val extra_tc_thms = map simplify_nested_tc extra_tcs
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   990
            val (r1,ind1) = fold simplify_tc tcs (r,ind)
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   991
            val r2 = Rules.FILTER_DISCH_ALL(not o USyntax.is_WFR) r1
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   992
        in loop(rst, nthms@extra_tc_thms, r2::R, ind1)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   993
        end
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   994
   val rules_tcs = ListPair.zip (Rules.CONJUNCTS rules1, TCs)
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   995
   val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1)
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   996
in
41164
6854e9a40edc avoid ML structure aliases (especially single-letter abbreviations);
wenzelm
parents: 40314
diff changeset
   997
  {induction = ind2, rules = Rules.LIST_CONJ rules2, nested_tcs = extras}
23150
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   998
end;
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
   999
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
  1000
073a65f0bc40 moved TFL files to canonical place;
wenzelm
parents:
diff changeset
  1001
end;