src/HOL/Tools/Function/mutual.ML
author wenzelm
Fri, 17 Jul 2009 21:32:58 +0200
changeset 32029 94b4a921c88d
parent 31775 2b04504fcb69
child 32149 ef59550a55d3
permissions -rw-r--r--
compare types directly -- no need to invoke Type.eq_type with empty environment;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31775
2b04504fcb69 uniformly capitialized names for subdirectories
haftmann
parents: 30906
diff changeset
     1
(*  Title:      HOL/Tools/Function/mutual.ML
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
     2
    Author:     Alexander Krauss, TU Muenchen
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
     3
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
     4
A package for general recursive function definitions.
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
     5
Tools for mutual recursive definitions.
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
     6
*)
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
     7
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
     8
signature FUNDEF_MUTUAL =
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
     9
sig
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    10
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    11
  val prepare_fundef_mutual : FundefCommon.fundef_config
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    12
                              -> string (* defname *)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    13
                              -> ((string * typ) * mixfix) list
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    14
                              -> term list
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    15
                              -> local_theory
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    16
                              -> ((thm (* goalstate *)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    17
                                   * (thm -> FundefCommon.fundef_result) (* proof continuation *)
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    18
                                  ) * local_theory)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    19
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    20
end
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    21
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    22
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    23
structure FundefMutual: FUNDEF_MUTUAL =
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    24
struct
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    25
21051
c49467a9c1e1 Switched function package to use the new package for inductive predicates.
krauss
parents: 20949
diff changeset
    26
open FundefLib
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    27
open FundefCommon
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    28
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
    29
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
    30
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    31
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    32
type qgar = string * (string * typ) list * term list * term list * term
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    33
23217
8eac3bda1063 name_of_fqgar: precise type;
wenzelm
parents: 23203
diff changeset
    34
fun name_of_fqgar ((f, _, _, _, _): qgar) = f
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    35
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    36
datatype mutual_part =
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    37
  MutualPart of 
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    38
   {
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
    39
    i : int,
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
    40
    i' : int,
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    41
    fvar : string * typ,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    42
    cargTs: typ list,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    43
    f_def: term,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    44
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    45
    f: term option,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    46
    f_defthm : thm option
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    47
   }
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    48
   
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    49
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    50
datatype mutual_info =
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    51
  Mutual of 
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    52
   { 
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
    53
    n : int,
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
    54
    n' : int,
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    55
    fsum_var : string * typ,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    56
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    57
    ST: typ,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    58
    RST: typ,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    59
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    60
    parts: mutual_part list,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    61
    fqgars: qgar list,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    62
    qglrs: ((string * typ) list * term list * term * term) list,
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    63
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    64
    fsum : term option
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    65
   }
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    66
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    67
fun mutual_induct_Pnames n =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    68
    if n < 5 then fst (chop n ["P","Q","R","S"])
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    69
    else map (fn i => "P" ^ string_of_int i) (1 upto n)
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    70
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    71
fun get_part fname =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    72
    the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname)
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
    73
                     
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    74
(* FIXME *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    75
fun mk_prod_abs e (t1, t2) =
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    76
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    77
      val bTs = rev (map snd e)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    78
      val T1 = fastype_of1 (bTs, t1)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    79
      val T2 = fastype_of1 (bTs, t2)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    80
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    81
      HOLogic.pair_const T1 T2 $ t1 $ t2
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    82
    end;
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    83
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
    84
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
    85
fun analyze_eqs ctxt defname fs eqs =
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    86
    let
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
    87
      val num = length fs
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    88
        val fnames = map fst fs
24170
33f055a0f3a1 more error handling
krauss
parents: 23819
diff changeset
    89
        val fqgars = map (split_def ctxt) eqs
30906
3c7a76e79898 simplify computation and consistency checks of argument counts in the input
krauss
parents: 30607
diff changeset
    90
        val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
3c7a76e79898 simplify computation and consistency checks of argument counts in the input
krauss
parents: 30607
diff changeset
    91
                       |> AList.lookup (op =) #> the
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    92
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    93
        fun curried_types (fname, fT) =
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    94
            let
30906
3c7a76e79898 simplify computation and consistency checks of argument counts in the input
krauss
parents: 30607
diff changeset
    95
              val (caTs, uaTs) = chop (arity_of fname) (binder_types fT)
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    96
            in
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    97
                (caTs, uaTs ---> body_type fT)
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
    98
            end
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
    99
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   100
        val (caTss, resultTs) = split_list (map curried_types fs)
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   101
        val argTs = map (foldr1 HOLogic.mk_prodT) caTss
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   102
32029
94b4a921c88d compare types directly -- no need to invoke Type.eq_type with empty environment;
wenzelm
parents: 31775
diff changeset
   103
        val dresultTs = distinct (op =) resultTs
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   104
        val n' = length dresultTs
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   105
25555
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   106
        val RST = BalancedTree.make (uncurry SumTree.mk_sumT) dresultTs
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   107
        val ST = BalancedTree.make (uncurry SumTree.mk_sumT) argTs
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   108
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   109
        val fsum_type = ST --> RST
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   110
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   111
        val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   112
        val fsum_var = (fsum_var_name, fsum_type)
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   113
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   114
        fun define (fvar as (n, T)) caTs resultT i =
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   115
            let
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   116
                val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *)
32029
94b4a921c88d compare types directly -- no need to invoke Type.eq_type with empty environment;
wenzelm
parents: 31775
diff changeset
   117
                val i' = find_index (fn Ta => Ta = resultT) dresultTs + 1 
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   118
25555
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   119
                val f_exp = SumTree.mk_proj RST n' i' (Free fsum_var $ SumTree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars))
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   120
                val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp)
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   121
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   122
                val rew = (n, fold_rev lambda vars f_exp)
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   123
            in
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   124
                (MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew)
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   125
            end
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   126
            
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   127
        val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num))
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   128
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   129
        fun convert_eqs (f, qs, gs, args, rhs) =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   130
            let
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   131
              val MutualPart {i, i', ...} = get_part f parts
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   132
            in
25555
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   133
              (qs, gs, SumTree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args),
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   134
               SumTree.mk_inj RST n' i' (replace_frees rews rhs)
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   135
                               |> Envir.beta_norm)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   136
            end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   137
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   138
        val qglrs = map convert_eqs fqgars
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   139
    in
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   140
        Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, 
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   141
                parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE}
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   142
    end
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   143
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   144
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   145
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   146
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   147
fun define_projections fixes mutual fsum lthy =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   148
    let
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   149
      fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   150
          let
21436
5313a4cc3823 LocalTheory.notes/defs: proper kind;
wenzelm
parents: 21319
diff changeset
   151
            val ((f, (_, f_defthm)), lthy') =
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28083
diff changeset
   152
              LocalTheory.define Thm.internalK ((Binding.name fname, mixfix),
1de908189869 cleaned up binding module and related code
haftmann
parents: 28083
diff changeset
   153
                                            ((Binding.name (fname ^ "_def"), []), Term.subst_bound (fsum, f_def)))
21436
5313a4cc3823 LocalTheory.notes/defs: proper kind;
wenzelm
parents: 21319
diff changeset
   154
                              lthy
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   155
          in
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   156
            (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def,
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   157
                         f=SOME f, f_defthm=SOME f_defthm },
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   158
             lthy')
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   159
          end
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   160
          
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   161
      val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   162
      val (parts', lthy') = fold_map def (parts ~~ fixes) lthy
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   163
    in
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   164
      (Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts',
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   165
                fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum },
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   166
       lthy')
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   167
    end
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   168
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   169
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   170
fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F =
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   171
    let
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   172
      val thy = ProofContext.theory_of ctxt
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   173
                
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   174
      val oqnames = map fst pre_qs
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20654
diff changeset
   175
      val (qs, ctxt') = Variable.variant_fixes oqnames ctxt
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   176
                        |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   177
                        
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   178
      fun inst t = subst_bounds (rev qs, t)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   179
      val gs = map inst pre_gs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   180
      val args = map inst pre_args
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   181
      val rhs = inst pre_rhs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   182
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   183
      val cqs = map (cterm_of thy) qs
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   184
      val ags = map (assume o cterm_of thy) gs
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   185
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   186
      val import = fold forall_elim cqs
24977
9f98751c9628 replaced (flip Thm.implies_elim) by Thm.elim_implies;
wenzelm
parents: 24171
diff changeset
   187
                   #> fold Thm.elim_implies ags
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   188
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   189
      val export = fold_rev (implies_intr o cprop_of) ags
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   190
                   #> fold_rev forall_intr_rename (oqnames ~~ cqs)
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   191
    in
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   192
      F ctxt (f, qs, gs, args, rhs) import export
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   193
    end
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   194
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   195
fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs) import (export : thm -> thm) sum_psimp_eq =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   196
    let
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   197
      val (MutualPart {f=SOME f, f_defthm=SOME f_def, ...}) = get_part fname parts
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   198
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   199
      val psimp = import sum_psimp_eq
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   200
      val (simp, restore_cond) = case cprems_of psimp of
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   201
                                   [] => (psimp, I)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   202
                                 | [cond] => (implies_elim psimp (assume cond), implies_intr cond)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   203
                                 | _ => sys_error "Too many conditions"
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   204
    in
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   205
      Goal.prove ctxt [] [] 
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   206
                 (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs))
24171
25381ce95316 Issue a warning, when "function" encounters variables occuring in function position,
krauss
parents: 24170
diff changeset
   207
                 (fn _ => (LocalDefs.unfold_tac ctxt all_orig_fdefs)
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   208
                          THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 28965
diff changeset
   209
                          THEN (simp_tac (local_simpset_of ctxt addsimps SumTree.proj_in_rules)) 1)
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   210
        |> restore_cond 
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   211
        |> export
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   212
    end
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   213
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   214
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   215
(* FIXME HACK *)
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   216
fun mk_applied_form ctxt caTs thm =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   217
    let
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   218
      val thy = ProofContext.theory_of ctxt
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   219
      val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   220
    in
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   221
      fold (fn x => fn thm => combination thm (reflexive x)) xs thm
26196
0a0c2752561e Use conversions instead of simplifier. tuned
krauss
parents: 25555
diff changeset
   222
           |> Conv.fconv_rule (Thm.beta_conversion true)
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   223
           |> fold_rev forall_intr xs
26653
60e0cf6bef89 Thm.forall_elim_var(s);
wenzelm
parents: 26529
diff changeset
   224
           |> Thm.forall_elim_vars 0
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   225
    end
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   226
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   227
23494
f985f9239e0d removed "sum_tools.ML" in favour of BalancedTree
krauss
parents: 23217
diff changeset
   228
fun mutual_induct_rules lthy induct all_f_defs (Mutual {n, ST, RST, parts, ...}) =
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   229
    let
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22622
diff changeset
   230
      val cert = cterm_of (ProofContext.theory_of lthy)
20949
f030835fd9e4 Induction rules have schematic variables again.
krauss
parents: 20878
diff changeset
   231
      val newPs = map2 (fn Pname => fn MutualPart {cargTs, ...} => 
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   232
                                       Free (Pname, cargTs ---> HOLogic.boolT))
20949
f030835fd9e4 Induction rules have schematic variables again.
krauss
parents: 20878
diff changeset
   233
                       (mutual_induct_Pnames (length parts))
f030835fd9e4 Induction rules have schematic variables again.
krauss
parents: 20878
diff changeset
   234
                       parts
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   235
                       
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   236
      fun mk_P (MutualPart {cargTs, ...}) P =
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   237
          let
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   238
            val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   239
            val atup = foldr1 HOLogic.mk_prod avars
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   240
          in
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   241
            tupled_lambda atup (list_comb (P, avars))
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   242
          end
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   243
          
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   244
      val Ps = map2 mk_P parts newPs
25555
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   245
      val case_exp = SumTree.mk_sumcases HOLogic.boolT Ps
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   246
                     
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   247
      val induct_inst =
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22622
diff changeset
   248
          forall_elim (cert case_exp) induct
25555
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   249
                      |> full_simplify SumTree.sumcase_split_ss
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   250
                      |> full_simplify (HOL_basic_ss addsimps all_f_defs)
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   251
          
25046
3d0137a59dcb Fixed variable naming in mutual induction rules
krauss
parents: 25016
diff changeset
   252
      fun project rule (MutualPart {cargTs, i, ...}) k =
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   253
          let
25046
3d0137a59dcb Fixed variable naming in mutual induction rules
krauss
parents: 25016
diff changeset
   254
            val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *)
25555
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   255
            val inj = SumTree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs)
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   256
          in
25046
3d0137a59dcb Fixed variable naming in mutual induction rules
krauss
parents: 25016
diff changeset
   257
            (rule
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22622
diff changeset
   258
              |> forall_elim (cert inj)
25555
224a40e39457 factored out handling of sum types again
krauss
parents: 25046
diff changeset
   259
              |> full_simplify SumTree.sumcase_split_ss
25046
3d0137a59dcb Fixed variable naming in mutual induction rules
krauss
parents: 25016
diff changeset
   260
              |> fold_rev (forall_intr o cert) (afs @ newPs),
3d0137a59dcb Fixed variable naming in mutual induction rules
krauss
parents: 25016
diff changeset
   261
             k + length cargTs)
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   262
          end
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   263
    in
25046
3d0137a59dcb Fixed variable naming in mutual induction rules
krauss
parents: 25016
diff changeset
   264
      fst (fold_map (project induct_inst) parts 0)
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   265
    end
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   266
    
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   267
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22622
diff changeset
   268
fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof =
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   269
    let
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   270
      val result = inner_cont proof
26529
03ad378ed5f0 Function package no longer overwrites theorems.
krauss
parents: 26196
diff changeset
   271
      val FundefResult {fs=[f], G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   272
                        termination,domintros} = result
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   273
                                                                                                               
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22623
diff changeset
   274
      val (all_f_defs, fs) = map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22623
diff changeset
   275
                                     (mk_applied_form lthy cargTs (symmetric f_def), f))
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22623
diff changeset
   276
                                 parts
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22623
diff changeset
   277
                             |> split_list
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   278
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   279
      val all_orig_fdefs = map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   280
                           
20878
384c5bb713b2 mk_partial_rules_mutual: expand result terms/thms;
wenzelm
parents: 20851
diff changeset
   281
      fun mk_mpsimp fqgar sum_psimp =
22497
1fe951654cee fixed problem with the construction of mutual simp rules
krauss
parents: 22323
diff changeset
   282
          in_context lthy fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   283
          
23189
4574ab8f3b21 simplified interfaces, some restructuring
krauss
parents: 22733
diff changeset
   284
      val rew_ss = HOL_basic_ss addsimps all_f_defs
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   285
      val mpsimps = map2 mk_mpsimp fqgars psimps
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   286
      val mtrsimps = map_option (map2 mk_mpsimp fqgars) trsimps
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22622
diff changeset
   287
      val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m
23189
4574ab8f3b21 simplified interfaces, some restructuring
krauss
parents: 22733
diff changeset
   288
      val mtermination = full_simplify rew_ss termination
4574ab8f3b21 simplified interfaces, some restructuring
krauss
parents: 22733
diff changeset
   289
      val mdomintros = map_option (map (full_simplify rew_ss)) domintros
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   290
    in
22733
0b14bb35be90 definition lookup via terms, not names. Methods "relation" and "lexicographic_order"
krauss
parents: 22623
diff changeset
   291
      FundefResult { fs=fs, G=G, R=R,
26529
03ad378ed5f0 Function package no longer overwrites theorems.
krauss
parents: 26196
diff changeset
   292
                     psimps=mpsimps, simple_pinducts=minducts,
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22622
diff changeset
   293
                     cases=cases, termination=mtermination,
23189
4574ab8f3b21 simplified interfaces, some restructuring
krauss
parents: 22733
diff changeset
   294
                     domintros=mdomintros,
22623
5fcee5b319a2 proper handling of morphisms
krauss
parents: 22622
diff changeset
   295
                     trsimps=mtrsimps}
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 19922
diff changeset
   296
    end
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   297
      
23189
4574ab8f3b21 simplified interfaces, some restructuring
krauss
parents: 22733
diff changeset
   298
fun prepare_fundef_mutual config defname fixes eqss lthy =
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   299
    let
22323
b8c227d8ca91 beta_eta_contract specification befor processing. These normal forms avoid unpleasant surprises later on.
krauss
parents: 22244
diff changeset
   300
      val mutual = analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   301
      val Mutual {fsum_var=(n, T), qglrs, ...} = mutual
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   302
          
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   303
      val ((fsum, goalstate, cont), lthy') =
23189
4574ab8f3b21 simplified interfaces, some restructuring
krauss
parents: 22733
diff changeset
   304
          FundefCore.prepare_fundef config defname [((n, T), NoSyn)] qglrs lthy
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   305
          
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   306
      val (mutual', lthy'') = define_projections fixes mutual fsum lthy'
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   307
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   308
      val mutual_cont = mk_partial_rules_mutual lthy'' cont mutual'
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   309
    in
23819
2040846d1bbe some interface cleanup
krauss
parents: 23528
diff changeset
   310
      ((goalstate, mutual_cont), lthy'')
22166
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   311
    end
0a50d4db234a * Preliminary implementation of tail recursion
krauss
parents: 21436
diff changeset
   312
21237
b803f9870e97 untabified
krauss
parents: 21051
diff changeset
   313
    
19770
be5c23ebe1eb HOL/Tools/function_package: Added support for mutual recursive definitions.
krauss
parents:
diff changeset
   314
end