| author | blanchet | 
| Thu, 29 Apr 2010 19:08:02 +0200 | |
| changeset 36568 | d495d2e1f0a6 | 
| parent 36323 | 655e2d74de3a | 
| child 36954 | ef698bd61057 | 
| permissions | -rw-r--r-- | 
| 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: 
29097diff
changeset | 2 | Author: Norbert Voelker, FernUni Hagen | 
| 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
29097diff
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. | 
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31177diff
changeset | 6 | Taken from HOL/Tools/primrec.ML | 
| 21541 
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: 
29006diff
changeset | 11 | val add_primrec: term list option -> term option -> | 
| 29581 | 12 | (binding * typ option * mixfix) list -> | 
| 13 | (binding * typ option * mixfix) list -> | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 14 | (Attrib.binding * term) list -> local_theory -> Proof.state | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 15 | val add_primrec_cmd: string list option -> string option -> | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 16 | (binding * string option * mixfix) list -> | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 17 | (binding * string option * mixfix) list -> | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 18 | (Attrib.binding * string) list -> local_theory -> Proof.state | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 19 | end; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 20 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 21 | structure NominalPrimrec : NOMINAL_PRIMREC = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 22 | struct | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 23 | |
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33766diff
changeset | 24 | open Datatype_Aux; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 25 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 26 | exception RecError of string; | 
| 
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 | 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: 
29006diff
changeset | 29 | fun primrec_eq_err lthy s eq = | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 30 | 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 | 31 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 32 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 33 | (* preprocessing of equations *) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 34 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 35 | fun unquantify t = | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 36 | let | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 37 | 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: 
29006diff
changeset | 38 | val body = strip_qnt_body "all" t; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 39 | 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: 
29006diff
changeset | 40 | (fn Free (v, _) => insert (op =) v | _ => I) body [])) | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 41 | 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: 
29006diff
changeset | 42 | |
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 43 | fun process_eqn lthy is_fixed spec rec_fns = | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 44 | let | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 45 | val eq = unquantify spec; | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 46 | val (lhs, rhs) = | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 47 | 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: 
29006diff
changeset | 48 | handle TERM _ => raise RecError "not a proper equation"; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 49 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 50 | val (recfun, args) = strip_comb lhs; | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 51 | 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: 
29006diff
changeset | 52 | else raise RecError "illegal head of function equation" | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 53 | | _ => raise RecError "illegal head of function equation"; | 
| 21541 
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 (ls', rest) = take_prefix is_Free args; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 56 | val (middle, rs') = take_suffix is_Free rest; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 57 | val rpos = length ls'; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 58 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 59 | val (constr, cargs') = if null middle then raise RecError "constructor missing" | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 60 | else strip_comb (hd middle); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 61 | val (cname, T) = dest_Const constr | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 62 | handle TERM _ => raise RecError "ill-formed constructor"; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 63 | val (tname, _) = dest_Type (body_type T) handle TYPE _ => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 64 | raise RecError "cannot determine datatype associated with function" | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 65 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 66 | val (ls, cargs, rs) = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 67 | (map dest_Free ls', map dest_Free cargs', map dest_Free rs') | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 68 | handle TERM _ => raise RecError "illegal argument in pattern"; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 69 | val lfrees = ls @ rs @ cargs; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 70 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 71 | fun check_vars _ [] = () | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 72 | | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars)) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 73 | in | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 74 | if length middle > 1 then | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 75 | raise RecError "more than one non-variable in pattern" | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 76 | else | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 77 | (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 78 | check_vars "extra variables on rhs: " | 
| 29265 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 wenzelm parents: 
29097diff
changeset | 79 | (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: 
29006diff
changeset | 80 | |> filter_out (is_fixed o fst)); | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 81 | case AList.lookup (op =) rec_fns fname of | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 82 | NONE => | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 83 | (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 84 | | SOME (_, rpos', eqns) => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 85 | if AList.defined (op =) eqns cname then | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 86 | raise RecError "constructor already occurred as pattern" | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 87 | else if rpos <> rpos' then | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 88 | raise RecError "position of recursive argument inconsistent" | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 89 | else | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 90 | AList.update (op =) | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 91 | (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns)) | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 92 | rec_fns) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 93 | end | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 94 | handle RecError s => primrec_eq_err lthy s spec; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 95 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 96 | val param_err = "Parameters must be the same for all recursive functions"; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 97 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 98 | fun process_fun lthy descr eqns (i, fname) (fnames, fnss) = | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 99 | let | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 100 | val (_, (tname, _, constrs)) = nth descr i; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 101 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 102 | (* substitute "fname ls x rs" by "y" for (x, (_, y)) in subs *) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 103 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 104 | fun subst [] t fs = (t, fs) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 105 | | subst subs (Abs (a, T, t)) fs = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 106 | fs | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 107 | |> subst subs t | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 108 | |-> (fn t' => pair (Abs (a, T, t'))) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 109 | | subst subs (t as (_ $ _)) fs = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 110 | let | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 111 | val (f, ts) = strip_comb t; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 112 | in | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 113 | if is_Free f | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 114 | 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 | 115 | let | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 116 | val (fname', _) = dest_Free f; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 117 | val (_, rpos, eqns') = the (AList.lookup (op =) eqns fname'); | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 118 | val (ls, rs'') = chop rpos ts | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 119 | val (x', rs) = case rs'' of | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 120 | x' :: rs => (x', rs) | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 121 |                   | [] => raise RecError ("not enough arguments in recursive application\n"
 | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 122 | ^ "of function " ^ quote fname' ^ " on rhs"); | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 123 | val rs' = (case eqns' of | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 124 | (_, (ls', _, rs', _, _)) :: _ => | 
| 23006 
c46abff9a7a0
Fixed bug in subst causing primrec functions returning functions
 berghofe parents: 
22728diff
changeset | 125 | let val (rs1, rs2) = chop (length rs') rs | 
| 
c46abff9a7a0
Fixed bug in subst causing primrec functions returning functions
 berghofe parents: 
22728diff
changeset | 126 | in | 
| 
c46abff9a7a0
Fixed bug in subst causing primrec functions returning functions
 berghofe parents: 
22728diff
changeset | 127 | if ls = map Free ls' andalso rs1 = map Free rs' then rs2 | 
| 
c46abff9a7a0
Fixed bug in subst causing primrec functions returning functions
 berghofe parents: 
22728diff
changeset | 128 | else raise RecError param_err | 
| 
c46abff9a7a0
Fixed bug in subst causing primrec functions returning functions
 berghofe parents: 
22728diff
changeset | 129 | end | 
| 
c46abff9a7a0
Fixed bug in subst causing primrec functions returning functions
 berghofe parents: 
22728diff
changeset | 130 |                   | _ => raise RecError ("no equations for " ^ quote fname'));
 | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 131 | val (x, xs) = strip_comb x' | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 132 | in case AList.lookup (op =) subs x | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 133 | of NONE => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 134 | fs | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 135 | |> fold_map (subst subs) ts | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 136 | |-> (fn ts' => pair (list_comb (f, ts'))) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 137 | | SOME (i', y) => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 138 | fs | 
| 23006 
c46abff9a7a0
Fixed bug in subst causing primrec functions returning functions
 berghofe parents: 
22728diff
changeset | 139 | |> fold_map (subst subs) (xs @ rs') | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 140 | ||> process_fun lthy descr eqns (i', fname') | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 141 | |-> (fn ts' => pair (list_comb (y, ts'))) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 142 | end | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 143 | else | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 144 | fs | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 145 | |> fold_map (subst subs) (f :: ts) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 146 | |-> (fn (f'::ts') => pair (list_comb (f', ts'))) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 147 | end | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 148 | | subst _ t fs = (t, fs); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 149 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 150 | (* translate rec equations into function arguments suitable for rec comb *) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 151 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 152 | fun trans eqns (cname, cargs) (fnames', fnss', fns) = | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 153 | (case AList.lookup (op =) eqns cname of | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 154 |           NONE => (warning ("No equation for constructor " ^ quote cname ^
 | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 155 | "\nin definition of function " ^ quote fname); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 156 |               (fnames', fnss', (Const (@{const_name undefined}, dummyT))::fns))
 | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 157 | | SOME (ls, cargs', rs, rhs, eq) => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 158 | let | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 159 | val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 160 | val rargs = map fst recs; | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 161 | val subs = map (rpair dummyT o fst) | 
| 29276 | 162 | (rev (Term.rename_wrt_term rhs rargs)); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 163 | 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: 
29006diff
changeset | 164 | (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: 
29006diff
changeset | 165 | handle RecError s => primrec_eq_err lthy s eq | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 166 | in (fnames'', fnss'', | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 167 | (list_abs_free (cargs' @ subs, rhs'))::fns) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 168 | end) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 169 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 170 | in (case AList.lookup (op =) fnames i of | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 171 | NONE => | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 172 | if exists (fn (_, v) => fname = v) fnames then | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 173 |           raise RecError ("inconsistent functions for datatype " ^ quote tname)
 | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 174 | else | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 175 | let | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 176 | val SOME (_, _, eqns' as (_, (ls, _, rs, _, _)) :: _) = | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 177 | AList.lookup (op =) eqns fname; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 178 | val (fnames', fnss', fns) = fold_rev (trans eqns') constrs | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 179 | ((i, fname)::fnames, fnss, []) | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 180 | in | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 181 | (fnames', (i, (fname, ls, rs, fns))::fnss') | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 182 | end | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 183 | | SOME fname' => | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 184 | if fname = fname' then (fnames, fnss) | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 185 |         else raise RecError ("inconsistent functions for datatype " ^ quote tname))
 | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 186 | end; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 187 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 188 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 189 | (* prepare functions needed for definitions *) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 190 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 191 | fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 192 | case AList.lookup (op =) fns i of | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 193 | NONE => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 194 | let | 
| 28524 | 195 |          val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined},
 | 
| 33317 | 196 | replicate (length cargs + length (filter is_rec_type cargs)) | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 197 | dummyT ---> HOLogic.unitT)) constrs; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 198 |          val _ = warning ("No function definition for datatype " ^ quote tname)
 | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 199 | in | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 200 | (dummy_fns @ fs, defs) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 201 | end | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 202 | | 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 | 203 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 204 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 205 | (* make definition *) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 206 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 207 | 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 | 208 | let | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 209 | val used = map fst (fold Term.add_frees fs []); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 210 | val x = (Name.variant used "x", dummyT); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 211 | val frees = ls @ x :: rs; | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 212 | val raw_rhs = list_abs_free (frees, | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 213 | list_comb (Const (rec_name, dummyT), fs @ [Free x])) | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 214 | val def_name = Thm.def_name (Long_Name.base_name fname); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 215 | val rhs = singleton (Syntax.check_terms ctxt) raw_rhs; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 216 | 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: 
29581diff
changeset | 217 | 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: 
29006diff
changeset | 218 | in | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 219 | ((var, ((Binding.name def_name, []), rhs)), | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 220 | 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: 
29006diff
changeset | 221 | end; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 222 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 223 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 224 | (* find datatypes which contain all datatypes in tnames' *) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 225 | |
| 31938 | 226 | fun find_dts (dt_info : NominalDatatype.nominal_datatype_info Symtab.table) _ [] = [] | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 227 | | find_dts dt_info tnames' (tname::tnames) = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 228 | (case Symtab.lookup dt_info tname of | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 229 | NONE => primrec_err (quote tname ^ " is not a nominal datatype") | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 230 | | SOME dt => | 
| 33038 | 231 | if subset (op =) (tnames', map (#1 o snd) (#descr dt)) then | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 232 | (tname, dt)::(find_dts dt_info tnames' tnames) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 233 | else find_dts dt_info tnames' tnames); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 234 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 235 | fun common_prefix eq ([], _) = [] | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 236 | | common_prefix eq (_, []) = [] | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 237 | | common_prefix eq (x :: xs, y :: ys) = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 238 | if eq (x, y) then x :: common_prefix eq (xs, ys) else []; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 239 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 240 | local | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 241 | |
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 242 | fun gen_primrec 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 | 243 | let | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 244 | val (fixes', spec) = fst (prep_spec (raw_fixes @ raw_params) raw_spec lthy); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 245 | val fixes = List.take (fixes', length raw_fixes); | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 246 | val (names_atts, spec') = split_list spec; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 247 | val eqns' = map unquantify spec' | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 248 | 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: 
29581diff
changeset | 249 | orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) spec' []; | 
| 31938 | 250 | val dt_info = NominalDatatype.get_nominal_datatypes (ProofContext.theory_of lthy); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 251 | val lsrs :: lsrss = maps (fn (_, (_, _, eqns)) => | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 252 | map (fn (_, (ls, _, rs, _, _)) => ls @ rs) eqns) eqns | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 253 | val _ = | 
| 33038 | 254 | (if forall (curry (eq_set (op =)) lsrs) lsrss andalso forall | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 255 | (fn (_, (_, _, (_, (ls, _, rs, _, _)) :: eqns)) => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 256 | forall (fn (_, (ls', _, rs', _, _)) => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 257 | ls = ls' andalso rs = rs') eqns | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 258 | | _ => true) eqns | 
| 22330 
00ca68f5ce29
Replaced "raise RecError" by "primrec_err" in function
 berghofe parents: 
22101diff
changeset | 259 | then () else primrec_err param_err); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 260 | val tnames = distinct (op =) (map (#1 o snd) eqns); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 261 | val dts = find_dts dt_info tnames tnames; | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 262 | val main_fns = | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 263 |       map (fn (tname, {index, ...}) =>
 | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 264 | (index, | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 265 | (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 | 266 | dts; | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 267 |     val {descr, rec_names, rec_rewrites, ...} =
 | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 268 | if null dts then | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 269 |         primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
 | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 270 | else snd (hd dts); | 
| 22434 
081a62852313
- Replaced fold by fold_rev to make sure that list of predicate
 berghofe parents: 
22330diff
changeset | 271 | 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: 
22330diff
changeset | 272 | 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: 
22330diff
changeset | 273 | dTs' @ dTs @ [dT]) cargs [])) constrs))) descr; | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 274 | 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 | 275 | 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: 
29006diff
changeset | 276 | val defs' = map (make_def lthy fixes fs) defs; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 277 | val names1 = map snd fnames; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 278 | val names2 = map fst eqns; | 
| 33038 | 279 | val _ = if eq_set (op =) (names1, names2) then () | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 280 |       else primrec_err ("functions " ^ commas_quote names2 ^
 | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 281 | "\nare not mutually recursive"); | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 282 | val (defs_thms, lthy') = lthy |> | 
| 33766 
c679f05600cd
adapted Local_Theory.define -- eliminated odd thm kind;
 wenzelm parents: 
33726diff
changeset | 283 | fold_map (apfst (snd o snd) oo Local_Theory.define 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: 
29581diff
changeset | 284 | val qualify = Binding.qualify false | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 285 | (space_implode "_" (map (Long_Name.base_name o #1) defs)); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 286 | val names_atts' = map (apfst qualify) names_atts; | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 287 | val cert = cterm_of (ProofContext.theory_of lthy'); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 288 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 289 | fun mk_idx eq = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 290 | let | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 291 | 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 | 292 | (Logic.strip_imp_concl eq)))); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 293 | val SOME i = AList.lookup op = (map swap fnames) name; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 294 | val SOME (_, _, constrs) = AList.lookup op = descr i; | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 295 | val SOME (_, _, eqns'') = AList.lookup op = eqns name; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 296 | val SOME (cname, (_, cargs, _, _, _)) = find_first | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 297 | (fn (_, (_, _, _, _, eq')) => eq = eq') eqns'' | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 298 | in (i, find_index (fn (cname', _) => cname = cname') constrs, cargs) end; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 299 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 300 | val rec_rewritess = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 301 | unflat (map (fn (_, (_, _, constrs)) => constrs) descr) rec_rewrites; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 302 | val fvars = rec_rewrites |> hd |> concl_of |> HOLogic.dest_Trueprop |> | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 303 | HOLogic.dest_eq |> fst |> strip_comb |> snd |> take_prefix is_Var |> fst; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 304 | val (pvars, ctxtvars) = List.partition | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 305 | (equal HOLogic.boolT o body_type o snd) | 
| 35098 
45dec8e27c4b
Fixed bug in code for guessing the name of the variable representing the freshness context.
 berghofe parents: 
33968diff
changeset | 306 | (subtract (op =) | 
| 
45dec8e27c4b
Fixed bug in code for guessing the name of the variable representing the freshness context.
 berghofe parents: 
33968diff
changeset | 307 | (Term.add_vars (concl_of (hd rec_rewrites)) []) | 
| 
45dec8e27c4b
Fixed bug in code for guessing the name of the variable representing the freshness context.
 berghofe parents: 
33968diff
changeset | 308 | (fold_rev (Term.add_vars o Logic.strip_assums_concl) | 
| 
45dec8e27c4b
Fixed bug in code for guessing the name of the variable representing the freshness context.
 berghofe parents: 
33968diff
changeset | 309 | (prems_of (hd rec_rewrites)) [])); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 310 | val cfs = defs' |> hd |> snd |> strip_comb |> snd |> | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 311 | curry (List.take o swap) (length fvars) |> map cert; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 312 | val invs' = (case invs of | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 313 | NONE => map (fn (i, _) => | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 314 |           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: 
29006diff
changeset | 315 | | SOME invs' => map (prep_term lthy') invs'); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 316 | val inst = (map cert fvars ~~ cfs) @ | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 317 | (map (cert o Var) pvars ~~ map cert invs') @ | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 318 | (case ctxtvars of | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 319 | [ctxtvar] => [(cert (Var ctxtvar), | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 320 | cert (the_default HOLogic.unit (Option.map (prep_term lthy') fctxt)))] | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 321 | | _ => []); | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 322 | val rec_rewrites' = map (fn eq => | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 323 | let | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 324 | val (i, j, cargs) = mk_idx eq | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 325 | val th = nth (nth rec_rewritess i) j; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 326 | val cargs' = th |> concl_of |> HOLogic.dest_Trueprop |> | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 327 | HOLogic.dest_eq |> fst |> strip_comb |> snd |> split_last |> snd |> | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 328 | strip_comb |> snd | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 329 | in (cargs, Logic.strip_imp_prems eq, | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 330 | Drule.cterm_instantiate (inst @ | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 331 | (map cert cargs' ~~ map (cert o Free) cargs)) th) | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 332 | end) eqns'; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 333 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 334 | 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 | 335 | val cprems = map cert prems; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 336 | val asms = map Thm.assume cprems; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 337 | val premss = map (fn (cargs, eprems, eqn) => | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 338 | map (fn t => list_all_free (cargs, Logic.list_implies (eprems, t))) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 339 | (List.drop (prems_of eqn, length prems))) rec_rewrites'; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 340 | val cpremss = map (map cert) premss; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 341 | val asmss = map (map Thm.assume) cpremss; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 342 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 343 | fun mk_eqn ((cargs, eprems, eqn), asms') = | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 344 | let | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 345 | val ceprems = map cert eprems; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 346 | val asms'' = map Thm.assume ceprems; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 347 | val ccargs = map (cert o Free) cargs; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 348 | val asms''' = map (fn th => implies_elim_list | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 349 | (forall_elim_list ccargs th) asms'') asms' | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 350 | in | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 351 | implies_elim_list eqn (asms @ asms''') |> | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 352 | implies_intr_list ceprems |> | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 353 | forall_intr_list ccargs | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 354 | end; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 355 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 356 | val rule_prems = cprems @ flat cpremss; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 357 | val rule = implies_intr_list rule_prems | 
| 23418 | 358 | (Conjunction.intr_balanced (map mk_eqn (rec_rewrites' ~~ asmss))); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 359 | |
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 360 | val goals = map (fn ((cargs, _, _), eqn) => | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 361 | (list_all_free (cargs, eqn), [])) (rec_rewrites' ~~ eqns'); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 362 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 363 | in | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 364 | lthy' |> | 
| 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 365 | Variable.add_fixes (map fst lsrs) |> snd |> | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
35098diff
changeset | 366 | Proof.theorem NONE | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 367 | (fn thss => fn goal_ctxt => | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 368 | let | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 369 | val simps = ProofContext.export goal_ctxt lthy' (flat thss); | 
| 33671 | 370 | val (simps', lthy'') = | 
| 371 | fold_map Local_Theory.note (names_atts' ~~ map single simps) lthy'; | |
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 372 | in | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 373 | lthy'' | 
| 33671 | 374 | |> Local_Theory.note ((qualify (Binding.name "simps"), | 
| 33056 
791a4655cae3
renamed "nitpick_const_xxx" attributes to "nitpick_xxx" and "nitpick_ind_intros" to "nitpick_intros"
 blanchet parents: 
33040diff
changeset | 375 | map (Attrib.internal o K) [Simplifier.simp_add, Nitpick_Simps.add]), | 
| 30253 | 376 | maps snd simps') | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 377 | |> snd | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 378 | end) | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 379 | [goals] |> | 
| 30510 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 wenzelm parents: 
30487diff
changeset | 380 | Proof.apply (Method.Basic (fn _ => RAW_METHOD (fn _ => | 
| 29097 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 berghofe parents: 
29006diff
changeset | 381 | rewrite_goals_tac defs_thms THEN | 
| 32196 | 382 | compose_tac (false, rule, length rule_prems) 1))) |> | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 383 | Seq.hd | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 384 | end; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 385 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 386 | in | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 387 | |
| 33726 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 388 | val add_primrec = gen_primrec Specification.check_spec (K I); | 
| 
0878aecbf119
eliminated slightly odd name space grouping -- now managed by Isar toplevel;
 wenzelm parents: 
33671diff
changeset | 389 | val add_primrec_cmd = gen_primrec Specification.read_spec Syntax.read_term; | 
| 21541 
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: 
29006diff
changeset | 391 | end; | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 392 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 393 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 394 | (* outer syntax *) | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 395 | |
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 396 | local structure P = OuterParse in | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 397 | |
| 24906 
557a9cd9370c
turned keywords invariant/freshness_context into reserved indentifiers;
 wenzelm parents: 
24867diff
changeset | 398 | val freshness_context = P.reserved "freshness_context"; | 
| 
557a9cd9370c
turned keywords invariant/freshness_context into reserved indentifiers;
 wenzelm parents: 
24867diff
changeset | 399 | val invariant = P.reserved "invariant"; | 
| 24867 | 400 | |
| 24906 
557a9cd9370c
turned keywords invariant/freshness_context into reserved indentifiers;
 wenzelm parents: 
24867diff
changeset | 401 | fun unless_flag scan = Scan.unless ((freshness_context || invariant) -- P.$$$ ":") scan; | 
| 
557a9cd9370c
turned keywords invariant/freshness_context into reserved indentifiers;
 wenzelm parents: 
24867diff
changeset | 402 | |
| 
557a9cd9370c
turned keywords invariant/freshness_context into reserved indentifiers;
 wenzelm parents: 
24867diff
changeset | 403 | val parser1 = (freshness_context -- P.$$$ ":") |-- unless_flag P.term >> SOME; | 
| 
557a9cd9370c
turned keywords invariant/freshness_context into reserved indentifiers;
 wenzelm parents: 
24867diff
changeset | 404 | val parser2 = (invariant -- P.$$$ ":") |-- | 
| 
557a9cd9370c
turned keywords invariant/freshness_context into reserved indentifiers;
 wenzelm parents: 
24867diff
changeset | 405 | (Scan.repeat1 (unless_flag P.term) >> SOME) -- Scan.optional parser1 NONE || | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 406 | (parser1 >> pair NONE); | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 407 | val options = | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 408 |   Scan.optional (P.$$$ "(" |-- P.!!! (parser2 --| P.$$$ ")")) (NONE, NONE);
 | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 409 | |
| 24867 | 410 | val _ = | 
| 30487 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 411 | OuterSyntax.local_theory_to_proof "nominal_primrec" | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 412 | "define primitive recursive functions on nominal datatypes" OuterKeyword.thy_goal | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 413 | (options -- P.fixes -- P.for_fixes -- SpecParse.where_alt_specs | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 414 | >> (fn ((((invs, fctxt), fixes), params), specs) => | 
| 
a14ff49d3083
simplified preparation and outer parsing of specification;
 wenzelm parents: 
30364diff
changeset | 415 | add_primrec_cmd invs fctxt fixes params specs)); | 
| 21541 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 416 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 417 | end; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 418 | |
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 419 | end; | 
| 
ea881fbe0489
Implemented new "nominal_primrec" command for defining
 berghofe parents: diff
changeset | 420 |