author | wenzelm |
Sat, 16 Oct 2010 11:34:46 +0100 | |
changeset 39855 | d4299b415879 |
parent 39777 | 4f8f08362bf7 |
permissions | -rw-r--r-- |
39777
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
1 |
(* Title: HOL/Tools/old_primrec.ML |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
2 |
Author: Norbert Voelker, FernUni Hagen |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
3 |
Author: Stefan Berghofer, TU Muenchen |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
4 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
5 |
Package for defining functions on datatypes by primitive recursion. |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
6 |
*) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
7 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
8 |
signature OLD_PRIMREC = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
9 |
sig |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
10 |
val unify_consts: theory -> term list -> term list -> term list * term list |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
11 |
val add_primrec: string -> ((bstring * string) * Attrib.src list) list |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
12 |
-> theory -> thm list * theory |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
13 |
val add_primrec_unchecked: string -> ((bstring * string) * Attrib.src list) list |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
14 |
-> theory -> thm list * theory |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
15 |
val add_primrec_i: string -> ((bstring * term) * attribute list) list |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
16 |
-> theory -> thm list * theory |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
17 |
val add_primrec_unchecked_i: string -> ((bstring * term) * attribute list) list |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
18 |
-> theory -> thm list * theory |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
19 |
end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
20 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
21 |
structure OldPrimrec : OLD_PRIMREC = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
22 |
struct |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
23 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
24 |
open Datatype_Aux; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
25 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
26 |
exception RecError of string; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
27 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
28 |
fun primrec_err s = error ("Primrec definition error:\n" ^ s); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
29 |
fun primrec_eq_err thy s eq = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
30 |
primrec_err (s ^ "\nin\n" ^ quote (Syntax.string_of_term_global thy eq)); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
31 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
32 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
33 |
(*the following code ensures that each recursive set always has the |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
34 |
same type in all introduction rules*) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
35 |
fun unify_consts thy cs intr_ts = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
36 |
(let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
37 |
fun varify t (i, ts) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
38 |
let val t' = map_types (Logic.incr_tvar (i + 1)) (snd (Type.varify_global [] t)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
39 |
in (maxidx_of_term t', t'::ts) end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
40 |
val (i, cs') = fold_rev varify cs (~1, []); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
41 |
val (i', intr_ts') = fold_rev varify intr_ts (i, []); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
42 |
val rec_consts = fold Term.add_consts cs' []; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
43 |
val intr_consts = fold Term.add_consts intr_ts' []; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
44 |
fun unify (cname, cT) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
45 |
let val consts = map snd (filter (fn (c, _) => c = cname) intr_consts) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
46 |
in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
47 |
val (env, _) = fold unify rec_consts (Vartab.empty, i'); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
48 |
val subst = Type.legacy_freeze o map_types (Envir.norm_type env) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
49 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
50 |
in (map subst cs', map subst intr_ts') |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
51 |
end) handle Type.TUNIFY => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
52 |
(warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts)); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
53 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
54 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
55 |
(* preprocessing of equations *) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
56 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
57 |
fun process_eqn thy eq rec_fns = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
58 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
59 |
val (lhs, rhs) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
60 |
if null (Term.add_vars eq []) then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
61 |
HOLogic.dest_eq (HOLogic.dest_Trueprop eq) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
62 |
handle TERM _ => raise RecError "not a proper equation" |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
63 |
else raise RecError "illegal schematic variable(s)"; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
64 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
65 |
val (recfun, args) = strip_comb lhs; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
66 |
val fnameT = dest_Const recfun handle TERM _ => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
67 |
raise RecError "function is not declared as constant in theory"; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
68 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
69 |
val (ls', rest) = take_prefix is_Free args; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
70 |
val (middle, rs') = take_suffix is_Free rest; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
71 |
val rpos = length ls'; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
72 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
73 |
val (constr, cargs') = if null middle then raise RecError "constructor missing" |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
74 |
else strip_comb (hd middle); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
75 |
val (cname, T) = dest_Const constr |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
76 |
handle TERM _ => raise RecError "ill-formed constructor"; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
77 |
val (tname, _) = dest_Type (body_type T) handle TYPE _ => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
78 |
raise RecError "cannot determine datatype associated with function" |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
79 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
80 |
val (ls, cargs, rs) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
81 |
(map dest_Free ls', map dest_Free cargs', map dest_Free rs') |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
82 |
handle TERM _ => raise RecError "illegal argument in pattern"; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
83 |
val lfrees = ls @ rs @ cargs; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
84 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
85 |
fun check_vars _ [] = () |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
86 |
| check_vars s vars = raise RecError (s ^ commas_quote (map fst vars)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
87 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
88 |
if length middle > 1 then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
89 |
raise RecError "more than one non-variable in pattern" |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
90 |
else |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
91 |
(check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
92 |
check_vars "extra variables on rhs: " |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
93 |
(subtract (op =) lfrees (map dest_Free (OldTerm.term_frees rhs))); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
94 |
case AList.lookup (op =) rec_fns fnameT of |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
95 |
NONE => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
96 |
(fnameT, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eq))]))::rec_fns |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
97 |
| SOME (_, rpos', eqns) => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
98 |
if AList.defined (op =) eqns cname then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
99 |
raise RecError "constructor already occurred as pattern" |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
100 |
else if rpos <> rpos' then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
101 |
raise RecError "position of recursive argument inconsistent" |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
102 |
else |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
103 |
AList.update (op =) (fnameT, (tname, rpos, (cname, (ls, cargs, rs, rhs, eq))::eqns)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
104 |
rec_fns) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
105 |
end |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
106 |
handle RecError s => primrec_eq_err thy s eq; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
107 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
108 |
fun process_fun thy descr rec_eqns (i, fnameT as (fname, _)) (fnameTs, fnss) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
109 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
110 |
val (_, (tname, _, constrs)) = List.nth (descr, i); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
111 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
112 |
(* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
113 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
114 |
fun subst [] t fs = (t, fs) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
115 |
| subst subs (Abs (a, T, t)) fs = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
116 |
fs |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
117 |
|> subst subs t |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
118 |
|-> (fn t' => pair (Abs (a, T, t'))) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
119 |
| subst subs (t as (_ $ _)) fs = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
120 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
121 |
val (f, ts) = strip_comb t; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
122 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
123 |
if is_Const f andalso member (op =) (map fst rec_eqns) (dest_Const f) then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
124 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
125 |
val fnameT' as (fname', _) = dest_Const f; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
126 |
val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT'); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
127 |
val ls = take rpos ts; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
128 |
val rest = drop rpos ts; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
129 |
val (x', rs) = (hd rest, tl rest) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
130 |
handle Empty => raise RecError ("not enough arguments\ |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
131 |
\ in recursive application\nof function " ^ quote fname' ^ " on rhs"); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
132 |
val (x, xs) = strip_comb x' |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
133 |
in case AList.lookup (op =) subs x |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
134 |
of NONE => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
135 |
fs |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
136 |
|> fold_map (subst subs) ts |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
137 |
|-> (fn ts' => pair (list_comb (f, ts'))) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
138 |
| SOME (i', y) => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
139 |
fs |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
140 |
|> fold_map (subst subs) (xs @ ls @ rs) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
141 |
||> process_fun thy descr rec_eqns (i', fnameT') |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
142 |
|-> (fn ts' => pair (list_comb (y, ts'))) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
143 |
end |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
144 |
else |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
145 |
fs |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
146 |
|> fold_map (subst subs) (f :: ts) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
147 |
|-> (fn (f'::ts') => pair (list_comb (f', ts'))) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
148 |
end |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
149 |
| subst _ t fs = (t, fs); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
150 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
151 |
(* translate rec equations into function arguments suitable for rec comb *) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
152 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
153 |
fun trans eqns (cname, cargs) (fnameTs', fnss', fns) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
154 |
(case AList.lookup (op =) eqns cname of |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
155 |
NONE => (warning ("No equation for constructor " ^ quote cname ^ |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
156 |
"\nin definition of function " ^ quote fname); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
157 |
(fnameTs', fnss', (Const ("HOL.undefined", dummyT))::fns)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
158 |
| SOME (ls, cargs', rs, rhs, eq) => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
159 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
160 |
val recs = filter (is_rec_type o snd) (cargs' ~~ cargs); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
161 |
val rargs = map fst recs; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
162 |
val subs = map (rpair dummyT o fst) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
163 |
(rev (Term.rename_wrt_term rhs rargs)); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
164 |
val (rhs', (fnameTs'', fnss'')) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
165 |
(subst (map (fn ((x, y), z) => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
166 |
(Free x, (body_index y, Free z))) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
167 |
(recs ~~ subs)) rhs (fnameTs', fnss')) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
168 |
handle RecError s => primrec_eq_err thy s eq |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
169 |
in (fnameTs'', fnss'', |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
170 |
(list_abs_free (cargs' @ subs @ ls @ rs, rhs'))::fns) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
171 |
end) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
172 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
173 |
in (case AList.lookup (op =) fnameTs i of |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
174 |
NONE => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
175 |
if exists (equal fnameT o snd) fnameTs then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
176 |
raise RecError ("inconsistent functions for datatype " ^ quote tname) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
177 |
else |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
178 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
179 |
val (_, _, eqns) = the (AList.lookup (op =) rec_eqns fnameT); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
180 |
val (fnameTs', fnss', fns) = fold_rev (trans eqns) constrs |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
181 |
((i, fnameT)::fnameTs, fnss, []) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
182 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
183 |
(fnameTs', (i, (fname, #1 (snd (hd eqns)), fns))::fnss') |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
184 |
end |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
185 |
| SOME fnameT' => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
186 |
if fnameT = fnameT' then (fnameTs, fnss) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
187 |
else raise RecError ("inconsistent functions for datatype " ^ quote tname)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
188 |
end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
189 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
190 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
191 |
(* prepare functions needed for definitions *) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
192 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
193 |
fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
194 |
case AList.lookup (op =) fns i of |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
195 |
NONE => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
196 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
197 |
val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined", |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
198 |
replicate (length cargs + length (filter is_rec_type cargs)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
199 |
dummyT ---> HOLogic.unitT)) constrs; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
200 |
val _ = warning ("No function definition for datatype " ^ quote tname) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
201 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
202 |
(dummy_fns @ fs, defs) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
203 |
end |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
204 |
| SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
205 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
206 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
207 |
(* make definition *) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
208 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
209 |
fun make_def thy fs (fname, ls, rec_name, tname) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
210 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
211 |
val rhs = fold_rev (fn T => fn t => Abs ("", T, t)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
212 |
((map snd ls) @ [dummyT]) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
213 |
(list_comb (Const (rec_name, dummyT), |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
214 |
fs @ map Bound (0 ::(length ls downto 1)))) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
215 |
val def_name = Long_Name.base_name fname ^ "_" ^ Long_Name.base_name tname ^ "_def"; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
216 |
val def_prop = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
217 |
singleton (Syntax.check_terms (ProofContext.init_global thy)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
218 |
(Logic.mk_equals (Const (fname, dummyT), rhs)); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
219 |
in (def_name, def_prop) end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
220 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
221 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
222 |
(* find datatypes which contain all datatypes in tnames' *) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
223 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
224 |
fun find_dts (dt_info : info Symtab.table) _ [] = [] |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
225 |
| find_dts dt_info tnames' (tname::tnames) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
226 |
(case Symtab.lookup dt_info tname of |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
227 |
NONE => primrec_err (quote tname ^ " is not a datatype") |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
228 |
| SOME dt => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
229 |
if subset (op =) (tnames', map (#1 o snd) (#descr dt)) then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
230 |
(tname, dt)::(find_dts dt_info tnames' tnames) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
231 |
else find_dts dt_info tnames' tnames); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
232 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
233 |
fun prepare_induct ({descr, induct, ...}: info) rec_eqns = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
234 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
235 |
fun constrs_of (_, (_, _, cs)) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
236 |
map (fn (cname:string, (_, cargs, _, _, _)) => (cname, map fst cargs)) cs; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
237 |
val params_of = these o AList.lookup (op =) (maps constrs_of rec_eqns); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
238 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
239 |
induct |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
240 |
|> Rule_Cases.rename_params (map params_of (maps (map #1 o #3 o #2) descr)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
241 |
|> Rule_Cases.save induct |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
242 |
end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
243 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
244 |
local |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
245 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
246 |
fun gen_primrec_i note def alt_name eqns_atts thy = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
247 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
248 |
val (eqns, atts) = split_list eqns_atts; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
249 |
val dt_info = Datatype_Data.get_all thy; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
250 |
val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
251 |
val tnames = distinct (op =) (map (#1 o snd) rec_eqns); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
252 |
val dts = find_dts dt_info tnames tnames; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
253 |
val main_fns = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
254 |
map (fn (tname, {index, ...}) => |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
255 |
(index, |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
256 |
(fst o the o find_first (fn f => (#1 o snd) f = tname)) rec_eqns)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
257 |
dts; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
258 |
val {descr, rec_names, rec_rewrites, ...} = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
259 |
if null dts then |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
260 |
primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive") |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
261 |
else snd (hd dts); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
262 |
val (fnameTs, fnss) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
263 |
fold_rev (process_fun thy descr rec_eqns) main_fns ([], []); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
264 |
val (fs, defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
265 |
val defs' = map (make_def thy fs) defs; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
266 |
val nameTs1 = map snd fnameTs; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
267 |
val nameTs2 = map fst rec_eqns; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
268 |
val _ = if eq_set (op =) (nameTs1, nameTs2) then () |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
269 |
else primrec_err ("functions " ^ commas_quote (map fst nameTs2) ^ |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
270 |
"\nare not mutually recursive"); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
271 |
val primrec_name = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
272 |
if alt_name = "" then (space_implode "_" (map (Long_Name.base_name o #1) defs)) else alt_name; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
273 |
val (defs_thms', thy') = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
274 |
thy |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
275 |
|> Sign.add_path primrec_name |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
276 |
|> fold_map def (map (fn (name, t) => ((name, []), t)) defs'); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
277 |
val rewrites = (map mk_meta_eq rec_rewrites) @ map snd defs_thms'; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
278 |
val simps = map (fn (_, t) => Goal.prove_global thy' [] [] t |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
279 |
(fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
280 |
val (simps', thy'') = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
281 |
thy' |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
282 |
|> fold_map note ((map fst eqns ~~ atts) ~~ map single simps); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
283 |
val simps'' = maps snd simps'; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
284 |
val lhss = map (Logic.varify_global o fst o Logic.dest_equals o snd) defs'; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
285 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
286 |
thy'' |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
287 |
|> note (("simps", |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
288 |
[Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]), simps'') |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
289 |
|> snd |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
290 |
|> Spec_Rules.add_global Spec_Rules.Equational (lhss, simps) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
291 |
|> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns]) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
292 |
|> snd |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
293 |
|> Sign.parent_path |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
294 |
|> pair simps'' |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
295 |
end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
296 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
297 |
fun gen_primrec note def alt_name eqns thy = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
298 |
let |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
299 |
val ((names, strings), srcss) = apfst split_list (split_list eqns); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
300 |
val atts = map (map (Attrib.attribute thy)) srcss; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
301 |
val eqn_ts = map (fn s => Syntax.read_prop_global thy s |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
302 |
handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s)) strings; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
303 |
val rec_ts = map (fn eq => head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop eq))) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
304 |
handle TERM _ => primrec_eq_err thy "not a proper equation" eq) eqn_ts; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
305 |
val (_, eqn_ts') = unify_consts thy rec_ts eqn_ts |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
306 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
307 |
gen_primrec_i note def alt_name (names ~~ eqn_ts' ~~ atts) thy |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
308 |
end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
309 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
310 |
fun thy_note ((name, atts), thms) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
311 |
Global_Theory.add_thmss [((Binding.name name, thms), atts)] #-> (fn [thms] => pair (name, thms)); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
312 |
fun thy_def false ((name, atts), t) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
313 |
Global_Theory.add_defs false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm)) |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
314 |
| thy_def true ((name, atts), t) = |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
315 |
Global_Theory.add_defs_unchecked false [((Binding.name name, t), atts)] #-> (fn [thm] => pair (name, thm)); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
316 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
317 |
in |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
318 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
319 |
val add_primrec = gen_primrec thy_note (thy_def false); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
320 |
val add_primrec_unchecked = gen_primrec thy_note (thy_def true); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
321 |
val add_primrec_i = gen_primrec_i thy_note (thy_def false); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
322 |
val add_primrec_unchecked_i = gen_primrec_i thy_note (thy_def true); |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
323 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
324 |
end; |
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
325 |
|
4f8f08362bf7
moved old_primrec source to nominal package, where it is still used
haftmann
parents:
diff
changeset
|
326 |
end; |