src/HOL/Nominal/nominal_primrec.ML
author wenzelm
Thu, 05 Mar 2009 12:08:00 +0100
changeset 30280 eb98b49ef835
parent 30253 63cae7fd7e64
child 30364 577edc39b501
permissions -rw-r--r--
renamed NameSpace.base to NameSpace.base_name; renamed NameSpace.map_base to NameSpace.map_base_name; eliminated alias Sign.base_name = NameSpace.base_name;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Nominal/nominal_primrec.ML
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 29097
diff changeset
     2
    Author:     Norbert Voelker, FernUni Hagen
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 29097
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
     4
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
     5
Package for defining functions on nominal datatypes by primitive recursion.
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
     6
Taken from HOL/Tools/primrec_package.ML
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
     7
*)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
     8
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
     9
signature NOMINAL_PRIMREC =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    10
sig
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    11
  val add_primrec: term list option -> term option ->
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29276
diff changeset
    12
    (binding * typ option * mixfix) list ->
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29276
diff changeset
    13
    (binding * typ option * mixfix) list ->
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    14
    (Attrib.binding * term) list -> local_theory -> Proof.state
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    15
end;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    16
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    17
structure NominalPrimrec : NOMINAL_PRIMREC =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    18
struct
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    19
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    20
open DatatypeAux;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    21
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    22
exception RecError of string;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    23
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    24
fun primrec_err s = error ("Nominal primrec definition error:\n" ^ s);
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    25
fun primrec_eq_err lthy s eq =
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    26
  primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term lthy eq));
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    27
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    28
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    29
(* preprocessing of equations *)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    30
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    31
fun unquantify t =
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    32
  let
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    33
    val (vs, Ts) = split_list (strip_qnt_vars "all" t);
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    34
    val body = strip_qnt_body "all" t;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    35
    val (vs', _) = Name.variants vs (Name.make_context (fold_aterms
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    36
      (fn Free (v, _) => insert (op =) v | _ => I) body []))
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    37
  in (curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body) end;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    38
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    39
fun process_eqn lthy is_fixed spec rec_fns = 
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    40
  let
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    41
    val eq = unquantify spec;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    42
    val (lhs, rhs) = 
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    43
      HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_imp_concl eq))
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    44
      handle TERM _ => raise RecError "not a proper equation";
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    45
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    46
    val (recfun, args) = strip_comb lhs;
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    47
    val fname = case recfun of Free (v, _) => if is_fixed v then v
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    48
          else raise RecError "illegal head of function equation"
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    49
      | _ => raise RecError "illegal head of function equation";
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    50
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    51
    val (ls', rest)  = take_prefix is_Free args;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    52
    val (middle, rs') = take_suffix is_Free rest;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    53
    val rpos = length ls';
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    54
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    55
    val (constr, cargs') = if null middle then raise RecError "constructor missing"
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    56
      else strip_comb (hd middle);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    57
    val (cname, T) = dest_Const constr
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    58
      handle TERM _ => raise RecError "ill-formed constructor";
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    59
    val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    60
      raise RecError "cannot determine datatype associated with function"
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    61
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    62
    val (ls, cargs, rs) =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    63
      (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    64
      handle TERM _ => raise RecError "illegal argument in pattern";
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    65
    val lfrees = ls @ rs @ cargs;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    66
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    67
    fun check_vars _ [] = ()
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    68
      | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    69
  in
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    70
    if length middle > 1 then 
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    71
      raise RecError "more than one non-variable in pattern"
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    72
    else
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    73
     (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    74
      check_vars "extra variables on rhs: "
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 29097
diff changeset
    75
        (map dest_Free (OldTerm.term_frees rhs) |> subtract (op =) lfrees
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    76
          |> filter_out (is_fixed o fst));
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    77
      case AList.lookup (op =) rec_fns fname of
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    78
        NONE =>
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    79
          (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    80
      | SOME (_, rpos', eqns) =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    81
          if AList.defined (op =) eqns cname then
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    82
            raise RecError "constructor already occurred as pattern"
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    83
          else if rpos <> rpos' then
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    84
            raise RecError "position of recursive argument inconsistent"
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    85
          else
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    86
            AList.update (op =)
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    87
              (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns))
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    88
              rec_fns)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    89
  end
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    90
  handle RecError s => primrec_eq_err lthy s spec;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    91
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    92
val param_err = "Parameters must be the same for all recursive functions";
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    93
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    94
fun process_fun lthy descr eqns (i, fname) (fnames, fnss) =
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    95
  let
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
    96
    val (_, (tname, _, constrs)) = nth descr i;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    97
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    98
    (* substitute "fname ls x rs" by "y" for (x, (_, y)) in subs *)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
    99
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   100
    fun subst [] t fs = (t, fs)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   101
      | subst subs (Abs (a, T, t)) fs =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   102
          fs
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   103
          |> subst subs t
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   104
          |-> (fn t' => pair (Abs (a, T, t')))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   105
      | subst subs (t as (_ $ _)) fs =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   106
          let
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   107
            val (f, ts) = strip_comb t;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   108
          in
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   109
            if is_Free f
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   110
              andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   111
              let
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   112
                val (fname', _) = dest_Free f;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   113
                val (_, rpos, eqns') = the (AList.lookup (op =) eqns fname');
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   114
                val (ls, rs'') = chop rpos ts
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   115
                val (x', rs) = case rs'' of
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   116
                    x' :: rs => (x', rs)
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   117
                  | [] => raise RecError ("not enough arguments in recursive application\n"
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   118
                      ^ "of function " ^ quote fname' ^ " on rhs");
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   119
                val rs' = (case eqns' of
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   120
                    (_, (ls', _, rs', _, _)) :: _ =>
23006
c46abff9a7a0 Fixed bug in subst causing primrec functions returning functions
berghofe
parents: 22728
diff changeset
   121
                      let val (rs1, rs2) = chop (length rs') rs
c46abff9a7a0 Fixed bug in subst causing primrec functions returning functions
berghofe
parents: 22728
diff changeset
   122
                      in
c46abff9a7a0 Fixed bug in subst causing primrec functions returning functions
berghofe
parents: 22728
diff changeset
   123
                        if ls = map Free ls' andalso rs1 = map Free rs' then rs2
c46abff9a7a0 Fixed bug in subst causing primrec functions returning functions
berghofe
parents: 22728
diff changeset
   124
                        else raise RecError param_err
c46abff9a7a0 Fixed bug in subst causing primrec functions returning functions
berghofe
parents: 22728
diff changeset
   125
                      end
c46abff9a7a0 Fixed bug in subst causing primrec functions returning functions
berghofe
parents: 22728
diff changeset
   126
                  | _ => raise RecError ("no equations for " ^ quote fname'));
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   127
                val (x, xs) = strip_comb x'
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   128
              in case AList.lookup (op =) subs x
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   129
               of NONE =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   130
                    fs
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   131
                    |> fold_map (subst subs) ts
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   132
                    |-> (fn ts' => pair (list_comb (f, ts')))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   133
                | SOME (i', y) =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   134
                    fs
23006
c46abff9a7a0 Fixed bug in subst causing primrec functions returning functions
berghofe
parents: 22728
diff changeset
   135
                    |> fold_map (subst subs) (xs @ rs')
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   136
                    ||> process_fun lthy descr eqns (i', fname')
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   137
                    |-> (fn ts' => pair (list_comb (y, ts')))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   138
              end
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   139
            else
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   140
              fs
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   141
              |> fold_map (subst subs) (f :: ts)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   142
              |-> (fn (f'::ts') => pair (list_comb (f', ts')))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   143
          end
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   144
      | subst _ t fs = (t, fs);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   145
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   146
    (* translate rec equations into function arguments suitable for rec comb *)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   147
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   148
    fun trans eqns (cname, cargs) (fnames', fnss', fns) =
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   149
      (case AList.lookup (op =) eqns cname of
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   150
          NONE => (warning ("No equation for constructor " ^ quote cname ^
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   151
            "\nin definition of function " ^ quote fname);
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   152
              (fnames', fnss', (Const (@{const_name undefined}, dummyT))::fns))
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   153
        | SOME (ls, cargs', rs, rhs, eq) =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   154
            let
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   155
              val recs = filter (is_rec_type o snd) (cargs' ~~ cargs);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   156
              val rargs = map fst recs;
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   157
              val subs = map (rpair dummyT o fst)
29276
94b1ffec9201 qualified Term.rename_wrt_term;
wenzelm
parents: 29265
diff changeset
   158
                (rev (Term.rename_wrt_term rhs rargs));
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   159
              val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   160
                (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   161
                  handle RecError s => primrec_eq_err lthy s eq
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   162
            in (fnames'', fnss'', 
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   163
                (list_abs_free (cargs' @ subs, rhs'))::fns)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   164
            end)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   165
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   166
  in (case AList.lookup (op =) fnames i of
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   167
      NONE =>
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   168
        if exists (fn (_, v) => fname = v) fnames then
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   169
          raise RecError ("inconsistent functions for datatype " ^ quote tname)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   170
        else
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   171
          let
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   172
            val SOME (_, _, eqns' as (_, (ls, _, rs, _, _)) :: _) =
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   173
              AList.lookup (op =) eqns fname;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   174
            val (fnames', fnss', fns) = fold_rev (trans eqns') constrs
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   175
              ((i, fname)::fnames, fnss, []) 
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   176
          in
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   177
            (fnames', (i, (fname, ls, rs, fns))::fnss')
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   178
          end
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   179
    | SOME fname' =>
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   180
        if fname = fname' then (fnames, fnss)
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   181
        else raise RecError ("inconsistent functions for datatype " ^ quote tname))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   182
  end;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   183
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   184
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   185
(* prepare functions needed for definitions *)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   186
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   187
fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   188
  case AList.lookup (op =) fns i of
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   189
     NONE =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   190
       let
28524
644b62cf678f arbitrary is undefined
haftmann
parents: 28083
diff changeset
   191
         val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined},
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   192
           replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   193
             dummyT ---> HOLogic.unitT)) constrs;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   194
         val _ = warning ("No function definition for datatype " ^ quote tname)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   195
       in
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   196
         (dummy_fns @ fs, defs)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   197
       end
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   198
   | SOME (fname, ls, rs, fs') => (fs' @ fs, (fname, ls, rs, rec_name, tname) :: defs);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   199
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   200
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   201
(* make definition *)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   202
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   203
fun make_def ctxt fixes fs (fname, ls, rs, rec_name, tname) =
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   204
  let
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   205
    val used = map fst (fold Term.add_frees fs []);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   206
    val x = (Name.variant used "x", dummyT);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   207
    val frees = ls @ x :: rs;
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   208
    val raw_rhs = list_abs_free (frees,
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   209
      list_comb (Const (rec_name, dummyT), fs @ [Free x]))
30280
eb98b49ef835 renamed NameSpace.base to NameSpace.base_name;
wenzelm
parents: 30253
diff changeset
   210
    val def_name = Thm.def_name (NameSpace.base_name fname);
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   211
    val rhs = singleton (Syntax.check_terms ctxt) raw_rhs;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   212
    val SOME var = get_first (fn ((b, _), mx) =>
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 29581
diff changeset
   213
      if Binding.name_of b = fname then SOME (b, mx) else NONE) fixes;
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   214
  in
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   215
    ((var, ((Binding.name def_name, []), rhs)),
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   216
     subst_bounds (rev (map Free frees), strip_abs_body rhs))
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   217
  end;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   218
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   219
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   220
(* find datatypes which contain all datatypes in tnames' *)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   221
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   222
fun find_dts (dt_info : NominalPackage.nominal_datatype_info Symtab.table) _ [] = []
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   223
  | find_dts dt_info tnames' (tname::tnames) =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   224
      (case Symtab.lookup dt_info tname of
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   225
          NONE => primrec_err (quote tname ^ " is not a nominal datatype")
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   226
        | SOME dt =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   227
            if tnames' subset (map (#1 o snd) (#descr dt)) then
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   228
              (tname, dt)::(find_dts dt_info tnames' tnames)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   229
            else find_dts dt_info tnames' tnames);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   230
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   231
fun common_prefix eq ([], _) = []
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   232
  | common_prefix eq (_, []) = []
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   233
  | common_prefix eq (x :: xs, y :: ys) =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   234
      if eq (x, y) then x :: common_prefix eq (xs, ys) else [];
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   235
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   236
local
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   237
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   238
fun prepare_spec prep_spec ctxt raw_fixes raw_spec =
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   239
  let
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   240
    val ((fixes, spec), _) = prep_spec
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   241
      raw_fixes (map (single o apsnd single) raw_spec) ctxt
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   242
  in (fixes, map (apsnd the_single) spec) end;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   243
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   244
fun gen_primrec set_group prep_spec prep_term invs fctxt raw_fixes raw_params raw_spec lthy =
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   245
  let
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   246
    val (fixes', spec) = prepare_spec prep_spec lthy (raw_fixes @ raw_params) raw_spec;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   247
    val fixes = List.take (fixes', length raw_fixes);
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   248
    val (names_atts, spec') = split_list spec;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   249
    val eqns' = map unquantify spec'
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   250
    val eqns = fold_rev (process_eqn lthy (fn v => Variable.is_fixed lthy v
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 29581
diff changeset
   251
      orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) spec' [];
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   252
    val dt_info = NominalPackage.get_nominal_datatypes (ProofContext.theory_of lthy);
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   253
    val lsrs :: lsrss = maps (fn (_, (_, _, eqns)) =>
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   254
      map (fn (_, (ls, _, rs, _, _)) => ls @ rs) eqns) eqns
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   255
    val _ =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   256
      (if forall (curry eq_set lsrs) lsrss andalso forall
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   257
         (fn (_, (_, _, (_, (ls, _, rs, _, _)) :: eqns)) =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   258
               forall (fn (_, (ls', _, rs', _, _)) =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   259
                 ls = ls' andalso rs = rs') eqns
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   260
           | _ => true) eqns
22330
00ca68f5ce29 Replaced "raise RecError" by "primrec_err" in function
berghofe
parents: 22101
diff changeset
   261
       then () else primrec_err param_err);
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   262
    val tnames = distinct (op =) (map (#1 o snd) eqns);
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   263
    val dts = find_dts dt_info tnames tnames;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   264
    val main_fns = 
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   265
      map (fn (tname, {index, ...}) =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   266
        (index, 
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   267
          (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns))
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   268
      dts;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   269
    val {descr, rec_names, rec_rewrites, ...} = 
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   270
      if null dts then
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   271
        primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   272
      else snd (hd dts);
22434
081a62852313 - Replaced fold by fold_rev to make sure that list of predicate
berghofe
parents: 22330
diff changeset
   273
    val descr = map (fn (i, (tname, args, constrs)) => (i, (tname, args,
081a62852313 - Replaced fold by fold_rev to make sure that list of predicate
berghofe
parents: 22330
diff changeset
   274
      map (fn (cname, cargs) => (cname, fold (fn (dTs, dT) => fn dTs' =>
081a62852313 - Replaced fold by fold_rev to make sure that list of predicate
berghofe
parents: 22330
diff changeset
   275
        dTs' @ dTs @ [dT]) cargs [])) constrs))) descr;
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   276
    val (fnames, fnss) = fold_rev (process_fun lthy descr eqns) main_fns ([], []);
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   277
    val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   278
    val defs' = map (make_def lthy fixes fs) defs;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   279
    val names1 = map snd fnames;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   280
    val names2 = map fst eqns;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   281
    val _ = if gen_eq_set (op =) (names1, names2) then ()
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   282
      else primrec_err ("functions " ^ commas_quote names2 ^
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   283
        "\nare not mutually recursive");
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   284
    val (defs_thms, lthy') = lthy |>
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   285
      set_group ? LocalTheory.set_group (serial_string ()) |>
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   286
      fold_map (apfst (snd o snd) oo
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   287
        LocalTheory.define Thm.definitionK o fst) defs';
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 29581
diff changeset
   288
    val qualify = Binding.qualify false
30280
eb98b49ef835 renamed NameSpace.base to NameSpace.base_name;
wenzelm
parents: 30253
diff changeset
   289
      (space_implode "_" (map (NameSpace.base_name o #1) defs));
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   290
    val names_atts' = map (apfst qualify) names_atts;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   291
    val cert = cterm_of (ProofContext.theory_of lthy');
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   292
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   293
    fun mk_idx eq =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   294
      let
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   295
        val Free (name, _) = head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   296
          (Logic.strip_imp_concl eq))));
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   297
        val SOME i = AList.lookup op = (map swap fnames) name;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   298
        val SOME (_, _, constrs) = AList.lookup op = descr i;
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   299
        val SOME (_, _, eqns'') = AList.lookup op = eqns name;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   300
        val SOME (cname, (_, cargs, _, _, _)) = find_first
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   301
          (fn (_, (_, _, _, _, eq')) => eq = eq') eqns''
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   302
      in (i, find_index (fn (cname', _) => cname = cname') constrs, cargs) end;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   303
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   304
    val rec_rewritess =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   305
      unflat (map (fn (_, (_, _, constrs)) => constrs) descr) rec_rewrites;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   306
    val fvars = rec_rewrites |> hd |> concl_of |> HOLogic.dest_Trueprop |>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   307
      HOLogic.dest_eq |> fst |> strip_comb |> snd |> take_prefix is_Var |> fst;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   308
    val (pvars, ctxtvars) = List.partition
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   309
      (equal HOLogic.boolT o body_type o snd)
22434
081a62852313 - Replaced fold by fold_rev to make sure that list of predicate
berghofe
parents: 22330
diff changeset
   310
      (fold_rev Term.add_vars (map Logic.strip_assums_concl
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   311
        (prems_of (hd rec_rewrites))) [] \\ map dest_Var fvars);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   312
    val cfs = defs' |> hd |> snd |> strip_comb |> snd |>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   313
      curry (List.take o swap) (length fvars) |> map cert;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   314
    val invs' = (case invs of
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   315
        NONE => map (fn (i, _) =>
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   316
          Abs ("x", fastype_of (snd (nth defs' i)), HOLogic.true_const)) descr
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   317
      | SOME invs' => map (prep_term lthy') invs');
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   318
    val inst = (map cert fvars ~~ cfs) @
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   319
      (map (cert o Var) pvars ~~ map cert invs') @
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   320
      (case ctxtvars of
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   321
         [ctxtvar] => [(cert (Var ctxtvar),
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   322
           cert (the_default HOLogic.unit (Option.map (prep_term lthy') fctxt)))]
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   323
       | _ => []);
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   324
    val rec_rewrites' = map (fn eq =>
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   325
      let
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   326
        val (i, j, cargs) = mk_idx eq
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   327
        val th = nth (nth rec_rewritess i) j;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   328
        val cargs' = th |> concl_of |> HOLogic.dest_Trueprop |>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   329
          HOLogic.dest_eq |> fst |> strip_comb |> snd |> split_last |> snd |>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   330
          strip_comb |> snd
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   331
      in (cargs, Logic.strip_imp_prems eq,
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   332
        Drule.cterm_instantiate (inst @
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   333
          (map cert cargs' ~~ map (cert o Free) cargs)) th)
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   334
      end) eqns';
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   335
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   336
    val prems = foldr1 (common_prefix op aconv) (map (prems_of o #3) rec_rewrites');
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   337
    val cprems = map cert prems;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   338
    val asms = map Thm.assume cprems;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   339
    val premss = map (fn (cargs, eprems, eqn) =>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   340
      map (fn t => list_all_free (cargs, Logic.list_implies (eprems, t)))
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   341
        (List.drop (prems_of eqn, length prems))) rec_rewrites';
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   342
    val cpremss = map (map cert) premss;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   343
    val asmss = map (map Thm.assume) cpremss;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   344
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   345
    fun mk_eqn ((cargs, eprems, eqn), asms') =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   346
      let
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   347
        val ceprems = map cert eprems;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   348
        val asms'' = map Thm.assume ceprems;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   349
        val ccargs = map (cert o Free) cargs;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   350
        val asms''' = map (fn th => implies_elim_list
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   351
          (forall_elim_list ccargs th) asms'') asms'
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   352
      in
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   353
        implies_elim_list eqn (asms @ asms''') |>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   354
        implies_intr_list ceprems |>
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   355
        forall_intr_list ccargs
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   356
      end;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   357
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   358
    val rule_prems = cprems @ flat cpremss;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   359
    val rule = implies_intr_list rule_prems
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23351
diff changeset
   360
      (Conjunction.intr_balanced (map mk_eqn (rec_rewrites' ~~ asmss)));
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   361
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   362
    val goals = map (fn ((cargs, _, _), eqn) =>
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   363
      (list_all_free (cargs, eqn), [])) (rec_rewrites' ~~ eqns');
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   364
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   365
  in
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   366
    lthy' |>
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   367
    Variable.add_fixes (map fst lsrs) |> snd |>
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   368
    Proof.theorem_i NONE
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   369
      (fn thss => fn goal_ctxt =>
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   370
         let
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   371
           val simps = ProofContext.export goal_ctxt lthy' (flat thss);
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   372
           val (simps', lthy'') = fold_map (LocalTheory.note Thm.theoremK)
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   373
             (names_atts' ~~ map single simps) lthy'
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   374
         in
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   375
           lthy''
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   376
           |> LocalTheory.note Thm.theoremK ((qualify (Binding.name "simps"),
30253
63cae7fd7e64 Fix parentheses.
blanchet
parents: 30251
diff changeset
   377
                map (Attrib.internal o K)
63cae7fd7e64 Fix parentheses.
blanchet
parents: 30251
diff changeset
   378
                    [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add]),
63cae7fd7e64 Fix parentheses.
blanchet
parents: 30251
diff changeset
   379
                maps snd simps')
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   380
           |> snd
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   381
         end)
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   382
      [goals] |>
23351
3702086a15a3 Method.Basic: include position;
wenzelm
parents: 23006
diff changeset
   383
    Proof.apply (Method.Basic (fn _ => Method.RAW_METHOD (fn _ =>
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   384
      rewrite_goals_tac defs_thms THEN
23351
3702086a15a3 Method.Basic: include position;
wenzelm
parents: 23006
diff changeset
   385
      compose_tac (false, rule, length rule_prems) 1), Position.none)) |>
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   386
    Seq.hd
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   387
  end;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   388
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   389
in
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   390
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   391
val add_primrec = gen_primrec false Specification.check_specification (K I);
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   392
val add_primrec_cmd = gen_primrec true Specification.read_specification Syntax.read_term;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   393
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   394
end;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   395
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   396
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   397
(* outer syntax *)
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   398
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   399
local structure P = OuterParse and K = OuterKeyword in
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   400
24906
557a9cd9370c turned keywords invariant/freshness_context into reserved indentifiers;
wenzelm
parents: 24867
diff changeset
   401
val freshness_context = P.reserved "freshness_context";
557a9cd9370c turned keywords invariant/freshness_context into reserved indentifiers;
wenzelm
parents: 24867
diff changeset
   402
val invariant = P.reserved "invariant";
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   403
24906
557a9cd9370c turned keywords invariant/freshness_context into reserved indentifiers;
wenzelm
parents: 24867
diff changeset
   404
fun unless_flag scan = Scan.unless ((freshness_context || invariant) -- P.$$$ ":") scan;
557a9cd9370c turned keywords invariant/freshness_context into reserved indentifiers;
wenzelm
parents: 24867
diff changeset
   405
557a9cd9370c turned keywords invariant/freshness_context into reserved indentifiers;
wenzelm
parents: 24867
diff changeset
   406
val parser1 = (freshness_context -- P.$$$ ":") |-- unless_flag P.term >> SOME;
557a9cd9370c turned keywords invariant/freshness_context into reserved indentifiers;
wenzelm
parents: 24867
diff changeset
   407
val parser2 = (invariant -- P.$$$ ":") |--
557a9cd9370c turned keywords invariant/freshness_context into reserved indentifiers;
wenzelm
parents: 24867
diff changeset
   408
    (Scan.repeat1 (unless_flag P.term) >> SOME) -- Scan.optional parser1 NONE ||
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   409
  (parser1 >> pair NONE);
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   410
val options =
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   411
  Scan.optional (P.$$$ "(" |-- P.!!!
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   412
    (parser2 --| P.$$$ ")")) (NONE, NONE);
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   413
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   414
fun pipe_error t = P.!!! (Scan.fail_with (K
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   415
  (cat_lines ["Equations must be separated by " ^ quote "|", quote t])));
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   416
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   417
val statement = SpecParse.opt_thm_name ":" -- P.prop --| Scan.ahead
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   418
  ((P.term :-- pipe_error) || Scan.succeed ("",""));
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   419
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   420
val statements = P.enum1 "|" statement;
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   421
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   422
val primrec_decl = P.opt_target -- options --
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   423
  P.fixes -- P.for_fixes --| P.$$$ "where" -- statements;
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   424
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24712
diff changeset
   425
val _ =
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   426
  OuterSyntax.command "nominal_primrec" "define primitive recursive functions on nominal datatypes" K.thy_goal
29097
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   427
    (primrec_decl >> (fn ((((opt_target, (invs, fctxt)), raw_fixes), raw_params), raw_spec) =>
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   428
      Toplevel.print o Toplevel.local_theory_to_proof opt_target
68245155eb58 Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents: 29006
diff changeset
   429
        (add_primrec_cmd invs fctxt raw_fixes raw_params raw_spec)));
21541
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   430
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   431
end;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   432
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   433
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   434
end;
ea881fbe0489 Implemented new "nominal_primrec" command for defining
berghofe
parents:
diff changeset
   435