Nominal/nominal_mutual.ML
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Thu, 13 Mar 2014 09:21:31 +0000
changeset 3229 b52e8651591f
parent 3227 35bb5b013f0e
child 3231 188826f1ccdb
permissions -rw-r--r--
updated to Isabelle changes
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
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
     7
    Joachim Breitner helped with the auxiliary graph
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
     8
    definitions (7 August 2012)
2665
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
Mutual recursive nominal function definitions.
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
*)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
    13
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
signature NOMINAL_FUNCTION_MUTUAL =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
sig
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
    16
  val prepare_nominal_function_mutual : Nominal_Function_Common.nominal_function_config
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
    -> string (* defname *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
    -> ((string * typ) * mixfix) list
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
    -> term list
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
    -> local_theory
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
    -> ((thm (* goalstate *)
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2821
diff changeset
    22
        * (thm -> Nominal_Function_Common.nominal_function_result) (* proof continuation *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
       ) * local_theory)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
end
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
structure Nominal_Function_Mutual: NOMINAL_FUNCTION_MUTUAL =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
struct
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
open Function_Lib
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
    30
open Function_Common
2819
4bd584ff4fab added an option for an invariant (at the moment only a stub)
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
    31
open Nominal_Function_Common
2665
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
type qgar = string * (string * typ) list * term list * term list * term
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
datatype mutual_part = MutualPart of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
 {i : int,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  i' : int,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  fvar : string * typ,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  cargTs: typ list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  f_def: term,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  f: term option,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  f_defthm : thm option}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
datatype mutual_info = Mutual of
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
 {n : int,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  n' : int,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  fsum_var : string * typ,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  ST: typ,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  RST: typ,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  parts: mutual_part list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  fqgars: qgar list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  qglrs: ((string * typ) list * term list * term * term) list,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  fsum : term option}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
fun mutual_induct_Pnames n =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  if n < 5 then fst (chop n ["P","Q","R","S"])
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  else map (fn i => "P" ^ string_of_int i) (1 upto n)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
fun get_part fname =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
(* FIXME *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
fun mk_prod_abs e (t1, t2) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
    val bTs = rev (map snd e)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
    val T1 = fastype_of1 (bTs, t1)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
    val T2 = fastype_of1 (bTs, t2)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
    HOLogic.pair_const T1 T2 $ t1 $ t2
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
fun analyze_eqs ctxt defname fs eqs =
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 num = length fs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
    val fqgars = map (split_def ctxt (K true)) eqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
    val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
      |> AList.lookup (op =) #> the
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
    fun curried_types (fname, fT) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
        val (caTs, uaTs) = chop (arity_of fname) (binder_types fT)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
        (caTs, uaTs ---> body_type fT)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
    val (caTss, resultTs) = split_list (map curried_types fs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
    val argTs = map (foldr1 HOLogic.mk_prodT) caTss
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
    val dresultTs = distinct (op =) resultTs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
    val n' = length dresultTs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
    95
    val RST = Balanced_Tree.make (uncurry Sum_Tree.mk_sumT) dresultTs
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
    96
    val ST = Balanced_Tree.make (uncurry Sum_Tree.mk_sumT) argTs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
    val fsum_type = ST --> RST
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
    val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
    val fsum_var = (fsum_var_name, fsum_type)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
    fun define (fvar as (n, _)) caTs resultT i =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
        val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
        val i' = find_index (fn Ta => Ta = resultT) dresultTs + 1
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   108
        val f_exp = Sum_Tree.mk_proj RST n' i' (Free fsum_var $ Sum_Tree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
        val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
        val rew = (n, fold_rev lambda vars f_exp)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
        (MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
    val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
    fun convert_eqs (f, qs, gs, args, rhs) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
        val MutualPart {i, i', ...} = get_part f parts
2781
542ff50555f5 updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   121
        val rhs' = rhs
542ff50555f5 updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents: 2745
diff changeset
   122
             |> map_aterms (fn t as Free (n, _) => the_default t (AList.lookup (op =) rews n) | t => t)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
      in
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   124
        (qs, gs, Sum_Tree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args),
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   125
         Envir.beta_norm (Sum_Tree.mk_inj RST n' i' rhs'))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
    val qglrs = map convert_eqs fqgars
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
    Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
      parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE}
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
fun define_projections fixes mutual fsum lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
    fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
        val ((f, (_, f_defthm)), lthy') =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
          Local_Theory.define
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
            ((Binding.name fname, mixfix),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
              ((Binding.conceal (Binding.name (fname ^ "_def")), []),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
              Term.subst_bound (fsum, f_def))) lthy
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
        (MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
           f=SOME f, f_defthm=SOME f_defthm },
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
         lthy')
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
    val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
    val (parts', lthy') = fold_map def (parts ~~ fixes) lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
    (Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts',
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
       fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum },
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
     lthy')
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
  let
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2983
diff changeset
   159
    val thy = Proof_Context.theory_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
    val oqnames = map fst pre_qs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
    val (qs, _) = Variable.variant_fixes oqnames ctxt
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
      |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
    fun inst t = subst_bounds (rev qs, t)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
    val gs = map inst pre_gs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
    val args = map inst pre_args
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
    val rhs = inst pre_rhs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
    val cqs = map (cterm_of thy) qs
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   171
    val (ags, ctxt') = fold_map Thm.assume_hyps (map (cterm_of thy) gs) ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
    val import = fold Thm.forall_elim cqs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
      #> fold Thm.elim_implies ags
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
    val export = fold_rev (Thm.implies_intr o cprop_of) ags
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
      #> fold_rev forall_intr_rename (oqnames ~~ cqs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  in
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   179
    F ctxt' (f, qs, gs, args, rhs) import export
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
  import (export : thm -> thm) sum_psimp_eq =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
    val (MutualPart {f=SOME f, ...}) = get_part fname parts
2974
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
   186
 
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
    val psimp = import sum_psimp_eq
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   188
    val ((simp, restore_cond), ctxt') =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
      case cprems_of psimp of
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   190
        [] => ((psimp, I), ctxt)
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   191
      | [cond] =>
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   192
          let val (asm, ctxt') = Thm.assume_hyps cond ctxt
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   193
          in ((Thm.implies_elim psimp asm, Thm.implies_intr cond), ctxt') end
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
      | _ => raise General.Fail "Too many conditions"
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  in
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   196
    Goal.prove ctxt' [] []
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
      (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs))
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   198
      (fn _ => (Local_Defs.unfold_tac ctxt' all_orig_fdefs)
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   199
         THEN EqSubst.eqsubst_tac ctxt' [0] [simp] 1
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   200
         THEN (simp_tac ctxt') 1)
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   201
    |> restore_cond
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   202
    |> export
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   203
  end
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   204
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   205
val inl_perm = @{lemma "x = Inl y ==> projl (permute p x) = permute p (projl x)" by simp}
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   206
val inr_perm = @{lemma "x = Inr y ==> projr (permute p x) = permute p (projr x)" by simp}
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   207
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   208
fun recover_mutual_eqvt eqvt_thm all_orig_fdefs parts ctxt (fname, _, _, args, _)
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   209
  import (export : thm -> thm) sum_psimp_eq =
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   210
  let
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   211
    val (MutualPart {f=SOME f, ...}) = get_part fname parts
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   212
    
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   213
    val psimp = import sum_psimp_eq
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   214
    val ((cond, simp, restore_cond), ctxt') =
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   215
      case cprems_of psimp of
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   216
        [] => (([], psimp, I), ctxt)
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   217
      | [cond] =>
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   218
          let val (asm, ctxt') = Thm.assume_hyps cond ctxt
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   219
          in (([asm], Thm.implies_elim psimp asm, Thm.implies_intr cond), ctxt') end
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   220
      | _ => raise General.Fail "Too many conditions"
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   221
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   222
    val ([p], ctxt'') = ctxt'
3204
b69c8660de14 fixed problem with not fresh enough permutation name in nominal_primrec
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3197
diff changeset
   223
      |> fold Variable.declare_term args  
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   224
      |> Variable.variant_fixes ["p"]
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   225
    val p = Free (p, @{typ perm})
3204
b69c8660de14 fixed problem with not fresh enough permutation name in nominal_primrec
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3197
diff changeset
   226
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   227
    val simpset =
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   228
      put_simpset HOL_basic_ss ctxt'' addsimps 
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   229
      @{thms permute_sum.simps[symmetric] Pair_eqvt[symmetric] sum.sel} @
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   230
      [(cond MRS eqvt_thm) RS @{thm sym}] @ 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   231
      [inl_perm, inr_perm, simp] 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   232
    val goal_lhs = mk_perm p (list_comb (f, args))
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   233
    val goal_rhs = list_comb (f, map (mk_perm p) args)
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   234
  in
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   235
    Goal.prove ctxt'' [] [] (HOLogic.Trueprop $ HOLogic.mk_eq (goal_lhs, goal_rhs))
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   236
      (fn _ => Local_Defs.unfold_tac ctxt'' all_orig_fdefs
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   237
         THEN (asm_full_simp_tac simpset 1))
3227
35bb5b013f0e updated with current AFP version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3226
diff changeset
   238
    |> singleton (Proof_Context.export ctxt'' ctxt)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
    |> restore_cond
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
    |> export
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
fun mk_applied_form ctxt caTs thm =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  let
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2983
diff changeset
   245
    val thy = Proof_Context.theory_of ctxt
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
    val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
    fold (fn x => fn thm => Thm.combination thm (Thm.reflexive x)) xs thm
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
    |> Conv.fconv_rule (Thm.beta_conversion true)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
    |> fold_rev Thm.forall_intr xs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
    |> Thm.forall_elim_vars 0
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   254
fun mutual_induct_rules ctxt induct all_f_defs (Mutual {n, ST, parts, ...}) =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  let
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   256
    val cert = cterm_of (Proof_Context.theory_of ctxt)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
    val newPs =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
      map2 (fn Pname => fn MutualPart {cargTs, ...} =>
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
          Free (Pname, cargTs ---> HOLogic.boolT))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
        (mutual_induct_Pnames (length parts)) parts
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
    fun mk_P (MutualPart {cargTs, ...}) P =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
        val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
        val atup = foldr1 HOLogic.mk_prod avars
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
        HOLogic.tupled_lambda atup (list_comb (P, avars))
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
    val Ps = map2 mk_P parts newPs
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   271
    val case_exp = Sum_Tree.mk_sumcases HOLogic.boolT Ps
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
    val induct_inst =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
      Thm.forall_elim (cert case_exp) induct
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   275
      |> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt)
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   276
      |> full_simplify (put_simpset HOL_basic_ss ctxt addsimps all_f_defs)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
    fun project rule (MutualPart {cargTs, i, ...}) k =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
      let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
        val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *)
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   281
        val inj = Sum_Tree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
      in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
        (rule
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
         |> Thm.forall_elim (cert inj)
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   285
         |> full_simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
         |> fold_rev (Thm.forall_intr o cert) (afs @ newPs),
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
         k + length cargTs)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
      end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  in
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
    fst (fold_map (project induct_inst) parts 0)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   293
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   294
fun forall_elim s (Const ("all", _) $ Abs (_, _, t)) = subst_bound (s, t)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   295
  | forall_elim _ t = t
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   296
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   297
val forall_elim_list = fold forall_elim
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   298
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   299
fun split_conj_thm th =
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   300
  (split_conj_thm (th RS conjunct1)) @ (split_conj_thm (th RS conjunct2)) handle THM _ => [th];
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   301
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   302
fun prove_eqvt ctxt fs argTss eqvts_thms induct_thms =
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   303
  let
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   304
    fun aux argTs s = argTs
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   305
      |> map (pair s)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   306
      |> Variable.variant_frees ctxt fs
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   307
    val argss' = map2 aux argTss (Name.invent (Variable.names_of ctxt) "" (length fs)) 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   308
    val argss = (map o map) Free argss'
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   309
    val arg_namess = (map o map) fst argss'
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   310
    val insts = (map o map) SOME arg_namess 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   311
   
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   312
    val ([p_name], ctxt') = Variable.variant_fixes ["p"] ctxt
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   313
    val p = Free (p_name, @{typ perm})
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   314
2983
Christian Urban <urbanc@in.tum.de>
parents: 2982
diff changeset
   315
    (* extracting the acc-premises from the induction theorems *)
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   316
    val acc_prems = 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   317
     map prop_of induct_thms
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   318
     |> map2 forall_elim_list argss 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   319
     |> map (strip_qnt_body "all")
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   320
     |> map (curry Logic.nth_prem 1)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   321
     |> map HOLogic.dest_Trueprop
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   322
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   323
    fun mk_goal acc_prem (f, args) = 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   324
      let
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   325
        val goal_lhs = mk_perm p (list_comb (f, args))
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   326
        val goal_rhs = list_comb (f, map (mk_perm p) args)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   327
      in
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   328
        HOLogic.mk_imp (acc_prem, HOLogic.mk_eq (goal_lhs, goal_rhs))
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   329
      end
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   330
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   331
    val goal = fold_conj_balanced (map2 mk_goal acc_prems (fs ~~ argss))
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   332
      |> HOLogic.mk_Trueprop
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   333
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   334
    val induct_thm = case induct_thms of
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   335
        [thm] => thm
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   336
          |> Drule.gen_all 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   337
          |> Thm.permute_prems 0 1
3226
780b7a2c50b6 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3219
diff changeset
   338
          |> (fn thm => atomize_rule ctxt' (length (prems_of thm) - 1) thm)
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   339
      | thms => thms
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   340
          |> map Drule.gen_all 
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   341
          |> map (Rule_Cases.add_consumes 1)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   342
          |> snd o Rule_Cases.strict_mutual_rule ctxt'
3226
780b7a2c50b6 updated to changes in Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3219
diff changeset
   343
          |> atomize_concl ctxt'
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   344
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   345
    fun tac thm = rtac (Drule.gen_all thm) THEN_ALL_NEW atac
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   346
  in
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   347
    Goal.prove ctxt' (flat arg_namess) [] goal
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   348
      (fn {context, ...} => HEADGOAL (DETERM o (rtac induct_thm) THEN' RANGE (map tac eqvts_thms)))
3045
d0ad264f8c4f updated to Isabelle 3 Nov; it includes a hack to work around a bug in the localised version of the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 2983
diff changeset
   349
    |> singleton (Proof_Context.export ctxt' ctxt)
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   350
    |> split_conj_thm
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   351
    |> map (fn th => th RS mp)
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   352
  end
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   353
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   354
fun mk_partial_rules_mutual ctxt inner_cont (m as Mutual {parts, fqgars, ...}) proof =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
    val result = inner_cont proof
3204
b69c8660de14 fixed problem with not fresh enough permutation name in nominal_primrec
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3197
diff changeset
   357
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2821
diff changeset
   358
    val NominalFunctionResult {G, R, cases, psimps, simple_pinducts=[simple_pinduct],
2974
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
   359
      termination, domintros, eqvts=[eqvt],...} = result
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
    val (all_f_defs, fs) =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
      map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   363
          (mk_applied_form ctxt cargTs (Thm.symmetric f_def), f))
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
      parts
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
      |> split_list
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
    val all_orig_fdefs =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
      map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   370
    val cargTss =
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   371
      map (fn MutualPart {f = SOME f, cargTs, ...} => cargTs) parts
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   372
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
    fun mk_mpsimp fqgar sum_psimp =
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   374
      in_context ctxt fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   376
    fun mk_meqvts fqgar sum_psimp =
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   377
      in_context ctxt fqgar (recover_mutual_eqvt eqvt all_orig_fdefs parts) sum_psimp
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   378
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   379
    val rew_simpset = put_simpset HOL_basic_ss ctxt addsimps all_f_defs
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
    val mpsimps = map2 mk_mpsimp fqgars psimps
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   381
    val minducts = mutual_induct_rules ctxt simple_pinduct all_f_defs m
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   382
    val mtermination = full_simplify rew_simpset termination
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   383
    val mdomintros = Option.map (map (full_simplify rew_simpset)) domintros
2981
c8acaded1777 temporary fix
Christian Urban <urbanc@in.tum.de>
parents: 2978
diff changeset
   384
    val meqvts = map2 mk_meqvts fqgars psimps
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   385
    val meqvt_funs = prove_eqvt ctxt fs cargTss meqvts minducts
2974
b95a2065aa10 generated the partial eqvt-theorem for functions
Christian Urban <urbanc@in.tum.de>
parents: 2973
diff changeset
   386
 in
2973
d1038e67923a added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
parents: 2821
diff changeset
   387
    NominalFunctionResult { fs=fs, G=G, R=R,
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
      psimps=mpsimps, simple_pinducts=minducts,
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
      cases=cases, termination=mtermination,
2982
4a00077c008f completed the eqvt-proofs for functions; they are stored under the name function_name.eqvt and added to the eqvt-list
Christian Urban <urbanc@in.tum.de>
parents: 2981
diff changeset
   390
      domintros=mdomintros, eqvts=meqvt_funs }
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
(* nominal *)
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   394
fun subst_all s (Q $ Abs(_, _, t)) = 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   395
  let
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   396
    val vs = map Free (Term.add_frees s [])
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   397
  in
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   398
    fold Logic.all vs (subst_bound (s, t))
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   399
  end
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   400
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   401
fun mk_comp_dummy t s = Const (@{const_name comp}, dummyT) $ t $ s
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   402
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   403
fun all v t = 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   404
  let
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   405
    val T = Term.fastype_of v
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   406
  in
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   407
    Logic.all_const T $ absdummy T (abstract_over (v, t)) 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   408
  end
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   409
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   410
(* nominal *)
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
fun prepare_nominal_function_mutual config defname fixes eqss lthy =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
  let
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
    val mutual as Mutual {fsum_var=(n, T), qglrs, ...} =
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
      analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   416
    val ((fsum, G, GIntro_thms, G_induct, goalstate, cont), lthy') =
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
      Nominal_Function_Core.prepare_nominal_function config defname [((n, T), NoSyn)] qglrs lthy
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   419
    val (mutual' as Mutual {n', parts, ST, RST, ...}, lthy'') = define_projections fixes mutual fsum lthy'
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   421
    (* defining the auxiliary graph *)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   422
    fun mk_cases (MutualPart {i', fvar as (n, T), ...}) =
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   423
      let
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   424
        val (tys, ty) = strip_type T
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   425
        val fun_var = Free (n ^ "_aux", HOLogic.mk_tupleT tys --> ty)
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   426
        val inj_fun = absdummy dummyT (Sum_Tree.mk_inj RST n' i' (Bound 0))
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   427
      in
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   428
        Syntax.check_term lthy'' (mk_comp_dummy inj_fun fun_var)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   429
      end
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   430
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   431
    val case_sum_exp = map mk_cases parts
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   432
      |> Sum_Tree.mk_sumcases RST 
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   433
   
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   434
    val (G_name, G_type) = dest_Free G 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   435
    val G_name_aux = G_name ^ "_aux"
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   436
    val subst = [(G, Free (G_name_aux, G_type))]
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   437
    val GIntros_aux = GIntro_thms
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   438
      |> map prop_of
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   439
      |> map (Term.subst_free subst)
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   440
      |> map (subst_all case_sum_exp)
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   441
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   442
    val ((G_aux, GIntro_aux_thms, _, G_aux_induct), lthy''') = 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   443
      Nominal_Function_Core.inductive_def ((Binding.name G_name_aux, G_type), NoSyn) GIntros_aux lthy''
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   444
3219
e5d9b6bca88c updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
   445
    val mutual_cont = mk_partial_rules_mutual lthy''' cont mutual'
e5d9b6bca88c updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3218
diff changeset
   446
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   447
    (* proof of equivalence between graph and auxiliary graph *)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   448
    val x = Var(("x", 0), ST)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   449
    val y = Var(("y", 1), RST)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   450
    val G_aux_prem = HOLogic.mk_Trueprop (G_aux $ x $ y)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   451
    val G_prem = HOLogic.mk_Trueprop (G $ x $ y)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   452
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   453
    fun mk_inj_goal  (MutualPart {i', ...}) =
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   454
      let
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   455
        val injs = Sum_Tree.mk_inj ST n' i' (Bound 0)
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   456
        val projs = y
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   457
          |> Sum_Tree.mk_proj RST n' i'
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   458
          |> Sum_Tree.mk_inj RST n' i'
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   459
      in
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   460
        Const (@{const_name "All"}, dummyT) $ absdummy dummyT
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   461
          (HOLogic.mk_imp (HOLogic.mk_eq(x, injs), HOLogic.mk_eq(projs, y)))
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   462
      end
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   463
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   464
    val goal_inj = Logic.mk_implies (G_aux_prem, 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   465
      HOLogic.mk_Trueprop (fold_conj (map mk_inj_goal parts)))
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   466
      |> all x |> all y
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   467
      |> Syntax.check_term lthy'''
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   468
    val goal_iff1 = Logic.mk_implies (G_aux_prem, G_prem)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   469
      |> all x |> all y
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   470
    val goal_iff2 = Logic.mk_implies (G_prem, G_aux_prem)
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   471
      |> all x |> all y
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   472
3229
b52e8651591f updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3227
diff changeset
   473
    val simp_thms = @{thms sum.sel sum.inject sum.case sum.distinct o_apply}
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   474
    val simpset0 = put_simpset HOL_basic_ss lthy''' addsimps simp_thms
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   475
    val simpset1 = put_simpset HOL_ss lthy''' addsimps simp_thms
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   476
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   477
    val inj_thm = Goal.prove lthy''' [] [] goal_inj 
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   478
      (K (HEADGOAL (DETERM o etac G_aux_induct THEN_ALL_NEW asm_simp_tac simpset1)))
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   479
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   480
    fun aux_tac thm = 
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   481
      rtac (Drule.gen_all thm) THEN_ALL_NEW (asm_full_simp_tac (simpset1 addsimps [inj_thm]))
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   482
    
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   483
    val iff1_thm = Goal.prove lthy''' [] [] goal_iff1 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   484
      (K (HEADGOAL (DETERM o etac G_aux_induct THEN' RANGE (map aux_tac GIntro_thms))))
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   485
      |> Drule.gen_all
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   486
    val iff2_thm = Goal.prove lthy''' [] [] goal_iff2 
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   487
      (K (HEADGOAL (DETERM o etac G_induct THEN' RANGE
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   488
        (map (aux_tac o simplify simpset0) GIntro_aux_thms))))
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   489
      |> Drule.gen_all
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   490
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   491
    val iff_thm = Goal.prove lthy''' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (G, G_aux)))
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   492
      (K (HEADGOAL (EVERY' ((map rtac @{thms ext ext iffI}) @ [etac iff2_thm, etac iff1_thm]))))
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   493
 
3218
89158f401b07 updated to simplifier changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 3204
diff changeset
   494
    val tac = HEADGOAL (simp_tac (put_simpset HOL_basic_ss lthy''' addsimps [iff_thm]))
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   495
    val goalstate' = 
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   496
      case (SINGLE tac) goalstate of
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   497
        NONE => error "auxiliary equivalence proof failed"
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   498
      | SOME st => st
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
  in
3197
25d11b449e92 definition of an auxiliary graph in nominal-primrec definitions
Christian Urban <urbanc@in.tum.de>
parents: 3045
diff changeset
   500
    ((goalstate', mutual_cont), lthy''')
2665
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
  end
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
16b5a67ee279 exported nominal function code to external file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
end