src/HOL/Tools/Function/partial_function.ML
author krauss
Mon, 13 Dec 2010 10:15:27 +0100
changeset 41365 d6379876ec4c
parent 40427 fe4a58419d46
child 41720 f6ab14e61604
permissions -rw-r--r--
eliminated dest_all_all_ctx
krauss@40353
     1
(*  Title:      HOL/Tools/Function/partial_function.ML
krauss@40353
     2
    Author:     Alexander Krauss, TU Muenchen
krauss@40353
     3
krauss@40353
     4
Partial function definitions based on least fixed points in ccpos.
krauss@40353
     5
*)
krauss@40353
     6
krauss@40353
     7
signature PARTIAL_FUNCTION =
krauss@40353
     8
sig
krauss@40353
     9
  val setup: theory -> theory
krauss@40353
    10
  val init: term -> term -> thm -> declaration
krauss@40353
    11
krauss@40353
    12
  val add_partial_function: string -> (binding * typ option * mixfix) list ->
krauss@40353
    13
    Attrib.binding * term -> local_theory -> local_theory
krauss@40353
    14
krauss@40353
    15
  val add_partial_function_cmd: string -> (binding * string option * mixfix) list ->
krauss@40353
    16
    Attrib.binding * string -> local_theory -> local_theory
krauss@40353
    17
end;
krauss@40353
    18
krauss@40353
    19
krauss@40353
    20
structure Partial_Function: PARTIAL_FUNCTION =
krauss@40353
    21
struct
krauss@40353
    22
krauss@40353
    23
(*** Context Data ***)
krauss@40353
    24
krauss@40353
    25
structure Modes = Generic_Data
krauss@40353
    26
(
krauss@40353
    27
  type T = ((term * term) * thm) Symtab.table;
krauss@40353
    28
  val empty = Symtab.empty;
krauss@40353
    29
  val extend = I;
krauss@40353
    30
  fun merge (a, b) = Symtab.merge (K true) (a, b);
krauss@40353
    31
)
krauss@40353
    32
krauss@40353
    33
fun init fixp mono fixp_eq phi =
krauss@40353
    34
  let
krauss@40353
    35
    val term = Morphism.term phi;
krauss@40353
    36
    val data' = ((term fixp, term mono), Morphism.thm phi fixp_eq);
krauss@40353
    37
    val mode = (* extract mode identifier from morphism prefix! *)
krauss@40353
    38
      Binding.prefix_of (Morphism.binding phi (Binding.empty))
krauss@40353
    39
      |> map fst |> space_implode ".";
krauss@40353
    40
  in
krauss@40353
    41
    if mode = "" then I
krauss@40353
    42
    else Modes.map (Symtab.update (mode, data'))
krauss@40353
    43
  end
krauss@40353
    44
krauss@40353
    45
val known_modes = Symtab.keys o Modes.get o Context.Proof;
krauss@40353
    46
val lookup_mode = Symtab.lookup o Modes.get o Context.Proof;
krauss@40353
    47
krauss@40353
    48
krauss@40353
    49
structure Mono_Rules = Named_Thms
krauss@40353
    50
(
krauss@40353
    51
  val name = "partial_function_mono";
krauss@40353
    52
  val description = "monotonicity rules for partial function definitions";
krauss@40353
    53
);
krauss@40353
    54
krauss@40353
    55
krauss@40353
    56
(*** Automated monotonicity proofs ***)
krauss@40353
    57
krauss@40353
    58
fun strip_cases ctac = ctac #> Seq.map snd;
krauss@40353
    59
krauss@40353
    60
(*rewrite conclusion with k-th assumtion*)
krauss@40353
    61
fun rewrite_with_asm_tac ctxt k =
krauss@40353
    62
  Subgoal.FOCUS (fn {context=ctxt', prems, ...} =>
krauss@40353
    63
    Local_Defs.unfold_tac ctxt' [nth prems k]) ctxt;
krauss@40353
    64
krauss@40353
    65
fun dest_case thy t =
krauss@40353
    66
  case strip_comb t of
krauss@40353
    67
    (Const (case_comb, _), args) =>
krauss@40353
    68
      (case Datatype.info_of_case thy case_comb of
krauss@40353
    69
         NONE => NONE
krauss@40353
    70
       | SOME {case_rewrites, ...} =>
krauss@40353
    71
           let
krauss@40353
    72
             val lhs = prop_of (hd case_rewrites)
krauss@40353
    73
               |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst;
krauss@40353
    74
             val arity = length (snd (strip_comb lhs));
krauss@40353
    75
             val conv = funpow (length args - arity) Conv.fun_conv
krauss@40353
    76
               (Conv.rewrs_conv (map mk_meta_eq case_rewrites));
krauss@40353
    77
           in
krauss@40353
    78
             SOME (nth args (arity - 1), conv)
krauss@40353
    79
           end)
krauss@40353
    80
  | _ => NONE;
krauss@40353
    81
krauss@40353
    82
(*split on case expressions*)
krauss@40353
    83
val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context=ctxt, ...} =>
krauss@40353
    84
  SUBGOAL (fn (t, i) => case t of
krauss@40353
    85
    _ $ (_ $ Abs (_, _, body)) =>
krauss@40353
    86
      (case dest_case (ProofContext.theory_of ctxt) body of
krauss@40353
    87
         NONE => no_tac
krauss@40353
    88
       | SOME (arg, conv) =>
krauss@40353
    89
           let open Conv in
krauss@40353
    90
              if not (null (loose_bnos arg)) then no_tac
krauss@40353
    91
              else ((DETERM o strip_cases o Induct.cases_tac ctxt false [[SOME arg]] NONE [])
krauss@40353
    92
                THEN_ALL_NEW (rewrite_with_asm_tac ctxt 0)
krauss@40353
    93
                THEN_ALL_NEW etac @{thm thin_rl}
krauss@40353
    94
                THEN_ALL_NEW (CONVERSION
krauss@40353
    95
                  (params_conv ~1 (fn ctxt' =>
krauss@40353
    96
                    arg_conv (arg_conv (abs_conv (K conv) ctxt'))) ctxt))) i
krauss@40353
    97
           end)
krauss@40353
    98
  | _ => no_tac) 1);
krauss@40353
    99
krauss@40353
   100
(*monotonicity proof: apply rules + split case expressions*)
krauss@40353
   101
fun mono_tac ctxt =
krauss@40353
   102
  K (Local_Defs.unfold_tac ctxt [@{thm curry_def}])
krauss@40353
   103
  THEN' (TRY o REPEAT_ALL_NEW
krauss@40353
   104
   (resolve_tac (Mono_Rules.get ctxt)
krauss@40353
   105
     ORELSE' split_cases_tac ctxt));
krauss@40353
   106
krauss@40353
   107
krauss@40353
   108
(*** Auxiliary functions ***)
krauss@40353
   109
krauss@40353
   110
(*positional instantiation with computed type substitution.
krauss@40353
   111
  internal version of  attribute "[of s t u]".*)
krauss@40353
   112
fun cterm_instantiate' cts thm =
krauss@40353
   113
  let
krauss@40353
   114
    val thy = Thm.theory_of_thm thm;
krauss@40353
   115
    val vs = rev (Term.add_vars (prop_of thm) [])
krauss@40353
   116
      |> map (Thm.cterm_of thy o Var);
krauss@40353
   117
  in
krauss@40353
   118
    cterm_instantiate (zip_options vs cts) thm
krauss@40353
   119
  end;
krauss@40353
   120
krauss@40353
   121
(*Returns t $ u, but instantiates the type of t to make the
krauss@40353
   122
application type correct*)
krauss@40353
   123
fun apply_inst ctxt t u =
krauss@40353
   124
  let
krauss@40353
   125
    val thy = ProofContext.theory_of ctxt;
krauss@40353
   126
    val T = domain_type (fastype_of t);
krauss@40353
   127
    val T' = fastype_of u;
krauss@40353
   128
    val subst = Type.typ_match (Sign.tsig_of thy) (T, T') Vartab.empty
krauss@40353
   129
      handle Type.TYPE_MATCH => raise TYPE ("apply_inst", [T, T'], [t, u])
krauss@40353
   130
  in
krauss@40353
   131
    map_types (Envir.norm_type subst) t $ u
krauss@40353
   132
  end;
krauss@40353
   133
krauss@40353
   134
fun head_conv cv ct =
krauss@40353
   135
  if can Thm.dest_comb ct then Conv.fun_conv (head_conv cv) ct else cv ct;
krauss@40353
   136
krauss@40353
   137
krauss@40353
   138
(*** currying transformation ***)
krauss@40353
   139
krauss@40353
   140
fun curry_const (A, B, C) =
krauss@40353
   141
  Const (@{const_name Product_Type.curry},
krauss@40353
   142
    [HOLogic.mk_prodT (A, B) --> C, A, B] ---> C);
krauss@40353
   143
krauss@40353
   144
fun mk_curry f =
krauss@40353
   145
  case fastype_of f of
krauss@40353
   146
    Type ("fun", [Type (_, [S, T]), U]) =>
krauss@40353
   147
      curry_const (S, T, U) $ f
krauss@40353
   148
  | T => raise TYPE ("mk_curry", [T], [f]);
krauss@40353
   149
krauss@40353
   150
(* iterated versions. Nonstandard left-nested tuples arise naturally
krauss@40353
   151
from "split o split o split"*)
krauss@40353
   152
fun curry_n arity = funpow (arity - 1) mk_curry;
krauss@40353
   153
fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_split;
krauss@40353
   154
krauss@40353
   155
val curry_uncurry_ss = HOL_basic_ss addsimps
krauss@40353
   156
  [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}]
krauss@40353
   157
krauss@40353
   158
krauss@40353
   159
(*** partial_function definition ***)
krauss@40353
   160
krauss@40353
   161
fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy =
krauss@40353
   162
  let
krauss@40353
   163
    val ((fixp, mono), fixp_eq) = the (lookup_mode lthy mode)
krauss@40353
   164
      handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".",
krauss@40353
   165
        "Known modes are " ^ commas_quote (known_modes lthy) ^ "."]);
krauss@40353
   166
krauss@40353
   167
    val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [eqn_raw] lthy;
krauss@41365
   168
    val ((_, plain_eqn), _) = Function_Lib.focus_term eqn lthy;
krauss@40353
   169
krauss@40353
   170
    val ((f_binding, fT), mixfix) = the_single fixes;
krauss@40353
   171
    val fname = Binding.name_of f_binding;
krauss@40353
   172
krauss@40353
   173
    val cert = cterm_of (ProofContext.theory_of lthy);
krauss@40353
   174
    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn);
krauss@40353
   175
    val (head, args) = strip_comb lhs;
krauss@40353
   176
    val F = fold_rev lambda (head :: args) rhs;
krauss@40353
   177
krauss@40353
   178
    val arity = length args;
krauss@40353
   179
    val (aTs, bTs) = chop arity (binder_types fT);
krauss@40353
   180
krauss@40353
   181
    val tupleT = foldl1 HOLogic.mk_prodT aTs;
krauss@40353
   182
    val fT_uc = tupleT :: bTs ---> body_type fT;
krauss@40353
   183
    val f_uc = Var ((fname, 0), fT_uc);
krauss@40353
   184
    val x_uc = Var (("x", 0), tupleT);
krauss@40353
   185
    val uncurry = lambda head (uncurry_n arity head);
krauss@40353
   186
    val curry = lambda f_uc (curry_n arity f_uc);
krauss@40353
   187
krauss@40353
   188
    val F_uc =
krauss@40353
   189
      lambda f_uc (uncurry_n arity (F $ curry_n arity f_uc));
krauss@40353
   190
krauss@40353
   191
    val mono_goal = apply_inst lthy mono (lambda f_uc (F_uc $ f_uc $ x_uc))
krauss@40353
   192
      |> HOLogic.mk_Trueprop
krauss@40353
   193
      |> Logic.all x_uc;
krauss@40353
   194
krauss@40353
   195
    val mono_thm = Goal.prove_internal [] (cert mono_goal)
krauss@40353
   196
        (K (mono_tac lthy 1))
krauss@40353
   197
      |> Thm.forall_elim (cert x_uc);
krauss@40353
   198
krauss@40353
   199
    val f_def_rhs = curry_n arity (apply_inst lthy fixp F_uc);
krauss@40353
   200
    val f_def_binding = Binding.conceal (Binding.name (Thm.def_name fname));
krauss@40353
   201
    val ((f, (_, f_def)), lthy') = Local_Theory.define
krauss@40353
   202
      ((f_binding, mixfix), ((f_def_binding, []), f_def_rhs)) lthy;
krauss@40353
   203
krauss@40353
   204
    val eqn = HOLogic.mk_eq (list_comb (f, args),
krauss@40353
   205
        Term.betapplys (F, f :: args))
krauss@40353
   206
      |> HOLogic.mk_Trueprop;
krauss@40353
   207
krauss@40353
   208
    val unfold =
krauss@40353
   209
      (cterm_instantiate' (map (SOME o cert) [uncurry, F, curry]) fixp_eq
krauss@40353
   210
        OF [mono_thm, f_def])
krauss@40353
   211
      |> Tactic.rule_by_tactic lthy (Simplifier.simp_tac curry_uncurry_ss 1);
krauss@40353
   212
krauss@40353
   213
    val rec_rule = let open Conv in
krauss@40353
   214
      Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ =>
krauss@40353
   215
        CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1
krauss@40353
   216
        THEN rtac @{thm refl} 1) end;
krauss@40353
   217
  in
krauss@40353
   218
    lthy'
krauss@40353
   219
    |> Local_Theory.note (eq_abinding, [rec_rule])
krauss@40353
   220
    |-> (fn (_, rec') =>
krauss@40421
   221
      Spec_Rules.add Spec_Rules.Equational ([f], rec')
krauss@40421
   222
      #> Local_Theory.note ((Binding.qualify true fname (Binding.name "simps"), []), rec') #> snd)
krauss@40353
   223
  end;
krauss@40353
   224
krauss@40353
   225
val add_partial_function = gen_add_partial_function Specification.check_spec;
krauss@40353
   226
val add_partial_function_cmd = gen_add_partial_function Specification.read_spec;
krauss@40353
   227
krauss@40353
   228
val mode = Parse.$$$ "(" |-- Parse.xname --| Parse.$$$ ")";
krauss@40353
   229
krauss@40353
   230
val _ = Outer_Syntax.local_theory
haftmann@40427
   231
  "partial_function" "define partial function" Keyword.thy_decl
krauss@40353
   232
  ((mode -- (Parse.fixes -- (Parse.where_ |-- Parse_Spec.spec)))
krauss@40353
   233
     >> (fn (mode, (fixes, spec)) => add_partial_function_cmd mode fixes spec));
krauss@40353
   234
krauss@40353
   235
krauss@40353
   236
val setup = Mono_Rules.setup;
krauss@40353
   237
krauss@40353
   238
end