Nominal/nominal_function.ML
author Christian Urban <urbanc@in.tum.de>
Wed, 16 Mar 2011 20:42:14 +0100
changeset 2745 34df2cffe259
parent 2665 16b5a67ee279
child 2752 9f44608ea28d
permissions -rw-r--r--
ported changes from function package....needs Isabelle 16 March or above
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 ->
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
    (Attrib.binding * term) list -> Function_Common.function_config ->
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 ->
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
    (Attrib.binding * string) list -> Function_Common.function_config ->
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 ->
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
    (Attrib.binding * term) list -> Function_Common.function_config ->
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 ->
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
    (Attrib.binding * string) list -> Function_Common.function_config ->
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
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
open Function_Common
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
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
    44
  [Simplifier.simp_add,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
   Code.add_default_eqn_attribute,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
   Nitpick_Simps.add]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
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
    49
  [Nitpick_Psimps.add]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
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
    52
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
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
    54
  simps lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
    val spec = post simps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
      |> map (apfst (apsnd (fn ats => moreatts @ ats)))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
      |> map (apfst (apfst extra_qualify))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
    val (saved_spec_simps, lthy) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
      fold_map Local_Theory.note spec lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
    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
    64
    val simps_by_f = sort saved_simps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
    fun add_for_f fname simps =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
      Local_Theory.note
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
        ((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
    69
      #> snd
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
    (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
    72
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
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
    76
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
    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
    78
    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
    79
    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
    80
    val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
    val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
    val defname = mk_defname fixes
2745
34df2cffe259 ported changes from function package....needs Isabelle 16 March or above
Christian Urban <urbanc@in.tum.de>
parents: 2665
diff changeset
    84
    val FunctionConfig {partials, default, ...} = config
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
    val _ =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
      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
    87
        "'function (default)'. Use 'partial_function'."
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
      else ()
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
    val ((goal_state, cont), lthy') =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
      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
    92
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
    fun afterqed [[proof]] lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
      let
2745
34df2cffe259 ported changes from function package....needs Isabelle 16 March or above
Christian Urban <urbanc@in.tum.de>
parents: 2665
diff changeset
    95
        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
    96
          termination, domintros, cases, ...} =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
          cont (Thm.close_derivation proof)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
        val fnames = map (fst o fst) fixes
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
        fun qualify n = Binding.name n
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
          |> Binding.qualify true defname
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
        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
   103
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
        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
   105
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
        val (((psimps', pinducts'), (_, [termination'])), lthy) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
          lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
          |> 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
   109
               "psimps" conceal_partial psimp_attribs psimps
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
          ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
                 [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
   112
                  Attrib.internal (K (Rule_Cases.consumes 1)),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
                  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
   114
          ||>> 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
   115
          ||> (snd o Local_Theory.note ((qualify "cases",
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
                 [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
   117
          ||> (case domintros of NONE => I | SOME thms => 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
                   Local_Theory.note ((qualify "domintros", []), thms) #> snd)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
        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
   121
          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
   122
          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
   123
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
        val _ =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
          if not is_external then ()
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
          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
   127
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
        (info, 
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
         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
   130
      end
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
    ((goal_state, afterqed), 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
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
   136
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
    val ((goal_state, afterqed), lthy') =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
      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
   139
    val pattern_thm =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
      case SINGLE (tac lthy') goal_state of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
        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
   142
      | SOME st => Goal.finish lthy' st
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
    lthy'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
    |> afterqed [[pattern_thm]]
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
val add_nominal_function =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  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
   150
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
   151
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
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
   153
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
    val ((goal_state, afterqed), lthy') =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
      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
   156
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
    lthy'
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
    |> 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
   159
    |> 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
   160
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
val nominal_function =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
  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
   164
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
   165
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
fun add_case_cong n thy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
    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
   169
      |> safe_mk_meta_eq
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
    Context.theory_map
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
      (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
   173
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
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
   176
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
(* setup *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
val setup =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
  Attrib.setup @{binding fundef_cong}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
    (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
   183
    "declaration of congruence rule for function definitions"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  #> setup_case_cong
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  #> Function_Relation.setup
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  #> Function_Common.Termination_Simps.setup
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
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
   189
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
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
   191
  |> the_single |> snd
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
(* outer syntax *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
(* nominal *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
val _ =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  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
   199
  Keyword.thy_goal
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  (function_parser default_config
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
     >> (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
   202
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
end