Nominal/nominal_function.ML
author Christian Urban <urbanc@in.tum.de>
Sun, 05 Jun 2011 21:14:23 +0100
changeset 2819 4bd584ff4fab
parent 2789 32979078bfe9
child 2821 c7d4bd9e89e0
permissions -rw-r--r--
added an option for an invariant (at the moment only a stub)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*  Nominal Mutual Functions
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
    Author:  Christian Urban
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
    heavily based on the code of Alexander Krauss
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
    (code forked on 14 January 2011)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
Main entry points to the nominal function package.
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
signature NOMINAL_FUNCTION =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
sig
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  include FUNCTION_DATA
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  val add_nominal_function: (binding * typ option * mixfix) list ->
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
    15
    (Attrib.binding * term) list ->  Nominal_Function_Common.nominal_function_config ->
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
    (Proof.context -> tactic) -> local_theory -> info * local_theory
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  val add_nominal_function_cmd: (binding * string option * mixfix) list ->
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
    19
    (Attrib.binding * string) list ->  Nominal_Function_Common.nominal_function_config ->
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
    (Proof.context -> tactic) -> local_theory -> info * local_theory
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  val nominal_function: (binding * typ option * mixfix) list ->
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
    23
    (Attrib.binding * term) list ->  Nominal_Function_Common.nominal_function_config ->
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
    local_theory -> Proof.state
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  val nominal_function_cmd: (binding * string option * mixfix) list ->
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
    27
    (Attrib.binding * string) list ->  Nominal_Function_Common.nominal_function_config ->
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
    local_theory -> Proof.state
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  val setup : theory -> theory
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  val get_congs : Proof.context -> thm list
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  val get_info : Proof.context -> term -> info
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
structure Nominal_Function : NOMINAL_FUNCTION =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
struct
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
open Function_Lib
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
    41
open Nominal_Function_Common
2752
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    42
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    43
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    44
(* Check for all sorts of errors in the input - nominal needs a different checking function *)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    45
fun nominal_check_defs ctxt fixes eqs =
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    46
  let
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    47
    val fnames = map (fst o fst) fixes
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    48
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    49
    fun check geq =
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    50
      let
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    51
        fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    52
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    53
        fun check_head fname =
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    54
          member (op =) fnames fname orelse
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    55
          input_error ("Illegal equation head. Expected " ^ commas_quote fnames)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    56
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    57
        val (fname, qs, gs, args, rhs) = split_def ctxt check_head geq
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    58
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    59
        val _ = length args > 0 orelse input_error "Function has no arguments:"
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    60
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    61
        fun add_bvs t is = add_loose_bnos (t, 0, is)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    62
        val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs []))
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    63
                    |> map (fst o nth (rev qs))
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    64
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    65
        val _ = forall (not o Term.exists_subterm
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    66
          (fn Free (n, _) => member (op =) fnames n | _ => false)) (gs @ args)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    67
          orelse input_error "Defined function may not occur in premises or arguments"
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    68
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    69
        val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    70
        val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    71
        val _ = null funvars orelse (warning (cat_lines
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    72
          ["Bound variable" ^ plural " " "s " funvars ^
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    73
          commas_quote (map fst funvars) ^ " occur" ^ plural "s" "" funvars ^
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    74
          " in function position.", "Misspelled constructor???"]); true)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    75
      in
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    76
        (fname, length args)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    77
      end
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    78
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    79
    val grouped_args = AList.group (op =) (map check eqs)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    80
    val _ = grouped_args
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    81
      |> map (fn (fname, ars) =>
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    82
        length (distinct (op =) ars) = 1
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    83
        orelse error ("Function " ^ quote fname ^
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    84
          " has different numbers of arguments in different equations"))
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    85
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    86
    val not_defined = subtract (op =) (map fst grouped_args) fnames
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    87
    val _ = null not_defined
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    88
      orelse error ("No defining equations for function" ^
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    89
        plural " " "s " not_defined ^ commas_quote not_defined)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    90
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    91
    fun check_sorts ((fname, fT), _) =
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    92
      Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    93
      orelse error (cat_lines
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    94
      ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    95
       Syntax.string_of_typ (Config.put show_sorts true ctxt) fT])
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    96
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    97
    val _ = map check_sorts fixes
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    98
  in
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
    99
    ()
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   100
  end
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   101
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   102
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   103
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
val simp_attribs = map (Attrib.internal o K)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  [Simplifier.simp_add,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
   Code.add_default_eqn_attribute,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
   Nitpick_Simps.add]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
val psimp_attribs = map (Attrib.internal o K)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
  [Nitpick_Psimps.add]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
fun add_simps fnames post sort extra_qualify label mod_binding moreatts
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  simps lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
    val spec = post simps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
      |> map (apfst (apsnd (fn ats => moreatts @ ats)))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
      |> map (apfst (apfst extra_qualify))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
    val (saved_spec_simps, lthy) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
      fold_map Local_Theory.note spec lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
    val saved_simps = maps snd saved_spec_simps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
    val simps_by_f = sort saved_simps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
    fun add_for_f fname simps =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
      Local_Theory.note
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
        ((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
      #> snd
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
    (saved_simps, fold2 add_for_f fnames simps_by_f lthy)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
fun prepare_nominal_function is_external prep default_constraint fixspec eqns config lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
    val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
    val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
    val fixes = map (apfst (apfst Binding.name_of)) fixes0;
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
    val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
2789
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2752
diff changeset
   142
    val (eqs, post, sort_cont, cnames) =  
32979078bfe9 functions involving if and case do not throw exceptions anymore; but eqvt_at assumption has now a precondition
Christian Urban <urbanc@in.tum.de>
parents: 2752
diff changeset
   143
      empty_preproc nominal_check_defs config ctxt' fixes spec (* nominal *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
    val defname = mk_defname fixes
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   146
    val NominalFunctionConfig {partials, default, ...} = config
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
    val _ =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
      if is_some default then Output.legacy_feature
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
        "'function (default)'. Use 'partial_function'."
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
      else ()
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
    val ((goal_state, cont), lthy') =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
      Nominal_Function_Mutual.prepare_nominal_function_mutual config defname fixes eqs lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
    fun afterqed [[proof]] lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
      let
2745
34df2cffe259 ported changes from function package....needs Isabelle 16 March or above
Christian Urban <urbanc@in.tum.de>
parents: 2665
diff changeset
   157
        val FunctionResult {fs, R, psimps, simple_pinducts,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
          termination, domintros, cases, ...} =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
          cont (Thm.close_derivation proof)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
        val fnames = map (fst o fst) fixes
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
        fun qualify n = Binding.name n
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
          |> Binding.qualify true defname
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
        val conceal_partial = if partials then I else Binding.conceal
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
        val addsmps = add_simps fnames post sort_cont
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
        val (((psimps', pinducts'), (_, [termination'])), lthy) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
          lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
          |> addsmps (conceal_partial o Binding.qualify false "partial")
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
               "psimps" conceal_partial psimp_attribs psimps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
          ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
                 [Attrib.internal (K (Rule_Cases.case_names cnames)),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
                  Attrib.internal (K (Rule_Cases.consumes 1)),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
                  Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
          ||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination])
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
          ||> (snd o Local_Theory.note ((qualify "cases",
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
                 [Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases]))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
          ||> (case domintros of NONE => I | SOME thms => 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
                   Local_Theory.note ((qualify "domintros", []), thms) #> snd)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
        val info = { add_simps=addsmps, case_names=cnames, psimps=psimps',
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
          pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
          fs=fs, R=R, defname=defname, is_partial=true }
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
        val _ =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
          if not is_external then ()
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
          else Specification.print_consts lthy (K false) (map fst fixes)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
        (info, 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
         lthy |> Local_Theory.declaration false (add_function_data o morph_function_data info))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
    ((goal_state, afterqed), lthy')
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
fun gen_add_nominal_function is_external prep default_constraint fixspec eqns config tac lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
    val ((goal_state, afterqed), lthy') =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
      prepare_nominal_function is_external prep default_constraint fixspec eqns config lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
    val pattern_thm =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
      case SINGLE (tac lthy') goal_state of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
        NONE => error "pattern completeness and compatibility proof failed"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
      | SOME st => Goal.finish lthy' st
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
    lthy'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
    |> afterqed [[pattern_thm]]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
val add_nominal_function =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  gen_add_nominal_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
val add_nominal_function_cmd = gen_add_nominal_function true Specification.read_spec "_::type"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
fun gen_nominal_function is_external prep default_constraint fixspec eqns config lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
    val ((goal_state, afterqed), lthy') =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
      prepare_nominal_function is_external prep default_constraint fixspec eqns config lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
    lthy'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
    |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (concl_of goal_state), [])]]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
    |> Proof.refine (Method.primitive_text (K goal_state)) |> Seq.hd
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
val nominal_function =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
  gen_nominal_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
val nominal_function_cmd = gen_nominal_function true Specification.read_spec "_::type"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
fun add_case_cong n thy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
    val cong = #case_cong (Datatype.the_info thy n)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
      |> safe_mk_meta_eq
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
    Context.theory_map
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
      (Function_Ctx_Tree.map_function_congs (Thm.add_thm cong)) thy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
2752
9f44608ea28d changed default preprocessor that does not catch variables only occuring on the right
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   240
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
(* setup *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
val setup =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  Attrib.setup @{binding fundef_cong}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
    (Attrib.add_del Function_Ctx_Tree.cong_add Function_Ctx_Tree.cong_del)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
    "declaration of congruence rule for function definitions"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  #> setup_case_cong
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  #> Function_Relation.setup
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
  #> Function_Common.Termination_Simps.setup
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
val get_congs = Function_Ctx_Tree.get_function_congs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  |> the_single |> snd
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
(* outer syntax *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   259
local
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   260
  val option_parser = Parse.group "option"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   261
    ((Parse.reserved "sequential" >> K Sequential)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   262
     || ((Parse.reserved "default" |-- Parse.term) >> Default)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   263
     || (Parse.reserved "domintros" >> K DomIntros)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   264
     || (Parse.reserved "no_partials" >> K No_Partials)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   265
     || ((Parse.reserved "invariant" |-- Parse.term) >> Invariant))
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   266
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   267
  fun config_parser default =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   268
    (Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Parse.list1 option_parser) --| Parse.$$$ ")") [])
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   269
     >> (fn opts => fold apply_opt opts default)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   270
in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   271
  fun nominal_function_parser default_cfg =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   272
      config_parser default_cfg -- Parse.fixes -- Parse_Spec.where_alt_specs
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   273
end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   274
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   275
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
val _ =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
  Outer_Syntax.local_theory_to_proof "nominal_primrec" "define general recursive nominal functions"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
  Keyword.thy_goal
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2789
diff changeset
   280
  (nominal_function_parser nominal_default_config
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
     >> (fn ((config, fixes), statements) => nominal_function_cmd fixes statements config))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
end