Nominal/nominal_function_common.ML
author Christian Urban <urbanc@in.tum.de>
Tue, 07 Jun 2011 08:52:59 +0100
changeset 2821 c7d4bd9e89e0
parent 2819 4bd584ff4fab
child 2822 23befefc6e73
permissions -rw-r--r--
fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*  Nominal Function Common
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
    Author: Christian Urban
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
     4
    heavily based on the code of Alexander Krauss
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
     5
    (code forked on 5 June 2011)
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
     6
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
Common definitions and other infrastructure for the function package.
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
*)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
signature FUNCTION_DATA =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
sig
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
type info =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
 {is_partial : bool,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  defname : string,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
    (* contains no logical entities: invariant under morphisms: *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  add_simps : (binding -> binding) -> string -> (binding -> binding) ->
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
    Attrib.src list -> thm list -> local_theory -> thm list * local_theory,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  case_names : string list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  fs : term list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  R : term,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  psimps: thm list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  pinducts: thm list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  simps : thm list option,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
  inducts : thm list option,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  termination: thm}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
structure Function_Data : FUNCTION_DATA =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
struct
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
type info =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
 {is_partial : bool,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  defname : string,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
    (* contains no logical entities: invariant under morphisms: *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  add_simps : (binding -> binding) -> string -> (binding -> binding) ->
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
    Attrib.src list -> thm list -> local_theory -> thm list * local_theory,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  case_names : string list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  fs : term list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  R : term,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  psimps: thm list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
  pinducts: thm list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  simps : thm list option,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  inducts : thm list option,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  termination: thm}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
structure Nominal_Function_Common =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
struct
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
open Function_Data
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
local open Function_Lib in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
(* Profiling *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
val profile = Unsynchronized.ref false;
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
fun PROFILE msg = if !profile then timeap_msg msg else I
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
val acc_const_name = @{const_name accp}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
fun mk_acc domT R =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R 
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
val function_name = suffix "C"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
val graph_name = suffix "_graph"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
val rel_name = suffix "_rel"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
val dom_name = suffix "_dom"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
(* Termination rules *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
structure TerminationRule = Generic_Data
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
(
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  type T = thm list
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  val empty = []
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  val extend = I
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  val merge = Thm.merge_thms
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
);
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
val get_termination_rules = TerminationRule.get
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
val store_termination_rule = TerminationRule.map o cons
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
(* Function definition result data *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
datatype function_result = FunctionResult of
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
 {fs: term list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
  G: term,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
  R: term,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  psimps : thm list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  simple_pinducts : thm list,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  cases : thm,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
  termination : thm,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  domintros : thm list option}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
fun morph_function_data ({add_simps, case_names, fs, R, psimps, pinducts,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
  simps, inducts, termination, defname, is_partial} : info) phi =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
    let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
      val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
      val name = Binding.name_of o Morphism.binding phi o Binding.name
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
    in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
      { add_simps = add_simps, case_names = case_names,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
        fs = map term fs, R = term R, psimps = fact psimps,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
        pinducts = fact pinducts, simps = Option.map fact simps,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
        inducts = Option.map fact inducts, termination = thm termination,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
        defname = name defname, is_partial=is_partial }
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
    end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
structure FunctionData = Generic_Data
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
(
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  type T = (term * info) Item_Net.T;
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst);
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  val extend = I;
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  fun merge tabs : T = Item_Net.merge tabs;
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
val get_function = FunctionData.get o Context.Proof;
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
fun lift_morphism thy f =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
    val term = Drule.term_rule thy f
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
    Morphism.thm_morphism f $> Morphism.term_morphism term
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
    $> Morphism.typ_morphism (Logic.type_map term)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
fun import_function_data t ctxt =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
    val thy = Proof_Context.theory_of ctxt
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
    val ct = cterm_of thy t
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
    val inst_morph = lift_morphism thy o Thm.instantiate
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
    fun match (trm, data) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
      SOME (morph_function_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
      handle Pattern.MATCH => NONE
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
    get_first match (Item_Net.retrieve (get_function ctxt) t)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
  end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
fun import_last_function ctxt =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  case Item_Net.content (get_function ctxt) of
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
    [] => NONE
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  | (t, data) :: _ =>
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
    let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
      val ([t'], ctxt') = Variable.import_terms true [t] ctxt
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
    in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
      import_function_data t' ctxt'
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
    end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
val all_function_data = Item_Net.content o get_function
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
fun add_function_data (data : info as {fs, termination, ...}) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  FunctionData.map (fold (fn f => Item_Net.update (f, data)) fs)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
  #> store_termination_rule termination
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
(* Simp rules for termination proofs *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
structure Termination_Simps = Named_Thms
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
(
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
  val name = "termination_simp"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  val description = "simplification rules for termination proofs"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
(* Default Termination Prover *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
structure TerminationProver = Generic_Data
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
(
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  type T = Proof.context -> tactic
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  val empty = (fn _ => error "Termination prover not configured")
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  val extend = I
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  fun merge (a, _) = a
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
val set_termination_prover = TerminationProver.put
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
val get_termination_prover = TerminationProver.get o Context.Proof
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
(* Configuration management *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
datatype nominal_function_opt
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  = Sequential
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
  | Default of string
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
  | DomIntros
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
  | No_Partials
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  | Invariant of string
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
datatype nominal_function_config = NominalFunctionConfig of
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
 {sequential: bool,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  default: string option,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  domintros: bool,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  partials: bool,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
  inv: string option}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
fun apply_opt Sequential (NominalFunctionConfig {sequential, default, domintros, partials, inv}) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
    NominalFunctionConfig 
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
      {sequential=true, default=default, domintros=domintros, partials=partials, inv=inv}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  | apply_opt (Default d) (NominalFunctionConfig {sequential, default, domintros, partials, inv}) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
    NominalFunctionConfig 
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
      {sequential=sequential, default=SOME d, domintros=domintros, partials=partials, inv=inv}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
  | apply_opt DomIntros (NominalFunctionConfig {sequential, default, domintros, partials, inv}) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
    NominalFunctionConfig 
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
      {sequential=sequential, default=default, domintros=true, partials=partials, inv=inv}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  | apply_opt No_Partials (NominalFunctionConfig {sequential, default, domintros, partials, inv}) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
    NominalFunctionConfig 
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
      {sequential=sequential, default=default, domintros=domintros, partials=false, inv=inv}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  | apply_opt (Invariant s) (NominalFunctionConfig {sequential, default, domintros, partials, inv}) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
    NominalFunctionConfig 
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
      {sequential=sequential, default=default, domintros=domintros, partials=partials, inv = SOME s}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
val nominal_default_config =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  NominalFunctionConfig { sequential=false, default=NONE,
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
    domintros=false, partials=true, inv=NONE}
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
(* Analyzing function equations *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
fun split_def ctxt check_head geq =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
    fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq]
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
    val qs = Term.strip_qnt_vars "all" geq
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
    val imp = Term.strip_qnt_body "all" geq
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
    val (gs, eq) = Logic.strip_horn imp
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
    val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
      handle TERM _ => error (input_error "Not an equation")
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
    val (head, args) = strip_comb f_args
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
    val fname = fst (dest_Free head) handle TERM _ => ""
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
    val _ = check_head fname
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
    (fname, qs, gs, args, rhs)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
(* Check for all sorts of errors in the input *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
fun check_defs ctxt fixes eqs =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
    val fnames = map (fst o fst) fixes
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
    fun check geq =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
      let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
        fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
        fun check_head fname =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
          member (op =) fnames fname orelse
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
          input_error ("Illegal equation head. Expected " ^ commas_quote fnames)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
        val (fname, qs, gs, args, rhs) = split_def ctxt check_head geq
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
        val _ = length args > 0 orelse input_error "Function has no arguments:"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
        fun add_bvs t is = add_loose_bnos (t, 0, is)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
        val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs []))
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
                    |> map (fst o nth (rev qs))
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
        val _ = null rvs orelse input_error
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
          ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs ^
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
           " occur" ^ plural "s" "" rvs ^ " on right hand side only:")
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
        val _ = forall (not o Term.exists_subterm
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
          (fn Free (n, _) => member (op =) fnames n | _ => false)) (gs @ args)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
          orelse input_error "Defined function may not occur in premises or arguments"
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
        val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
        val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
        val _ = null funvars orelse (warning (cat_lines
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
          ["Bound variable" ^ plural " " "s " funvars ^
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
          commas_quote (map fst funvars) ^ " occur" ^ plural "s" "" funvars ^
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
          " in function position.", "Misspelled constructor???"]); true)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
      in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
        (fname, length args)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
      end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
    val grouped_args = AList.group (op =) (map check eqs)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
    val _ = grouped_args
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
      |> map (fn (fname, ars) =>
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
        length (distinct (op =) ars) = 1
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
        orelse error ("Function " ^ quote fname ^
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
          " has different numbers of arguments in different equations"))
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
    val not_defined = subtract (op =) (map fst grouped_args) fnames
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
    val _ = null not_defined
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
      orelse error ("No defining equations for function" ^
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
        plural " " "s " not_defined ^ commas_quote not_defined)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
    fun check_sorts ((fname, fT), _) =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
      Sorts.of_sort (Sign.classes_of (Proof_Context.theory_of ctxt)) (fT, HOLogic.typeS)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
      orelse error (cat_lines
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
      ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
       Syntax.string_of_typ (Config.put show_sorts true ctxt) fT])
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
    val _ = map check_sorts fixes
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
  in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
    ()
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
  end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
(* Preprocessors *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
type fixes = ((string * typ) * mixfix) list
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
type 'a spec = (Attrib.binding * 'a list) list
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
type preproc = nominal_function_config -> Proof.context -> fixes -> term spec ->
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
val fname_of = fst o dest_Free o fst o strip_comb o fst o HOLogic.dest_eq o
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
  HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  | mk_case_names _ n 0 = []
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
  | mk_case_names _ n 1 = [n]
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
  | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
fun empty_preproc check _ ctxt fixes spec =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
  let
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
    val (bnds, tss) = split_list spec
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
    val ts = flat tss
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
    val _ = check ctxt fixes ts
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
    val fnames = map (fst o fst) fixes
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
    val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
    fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) 
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
      (indices ~~ xs) |> map (map snd)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
    (* using theorem names for case name currently disabled *)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
    val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
    (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
structure Preprocessor = Generic_Data
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
(
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
  type T = preproc
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  val empty : T = empty_preproc check_defs
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
  val extend = I
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
  fun merge (a, _) = a
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
)
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
val get_preproc = Preprocessor.get o Context.Proof
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
val set_preproc = Preprocessor.map o K
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
local
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
  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:
diff changeset
   349
    ((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:
diff changeset
   350
     || ((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:
diff changeset
   351
     || (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:
diff changeset
   352
     || (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:
diff changeset
   353
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
  fun config_parser default =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
    (Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Parse.list1 option_parser) --| Parse.$$$ ")") [])
2821
c7d4bd9e89e0 fixed problem with earlier commit about nominal_function_common; added facility for specifying an invariant - added a definition of frees_set which need a finiteness invariant
Christian Urban <urbanc@in.tum.de>
parents: 2819
diff changeset
   356
     >> (fn opts => fold Nominal_Function_Common.apply_opt opts default)
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
in
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
  fun function_parser default_cfg =
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
      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:
diff changeset
   360
end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
end
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
end