author | blanchet |
Thu, 26 Sep 2013 10:57:39 +0200 | |
changeset 53910 | 2c5055a3583d |
parent 53909 | 7c10e75e62b3 |
child 53918 | 0fc622be0185 |
permissions | -rw-r--r-- |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
1 |
(* Title: HOL/BNF/Tools/bnf_fp_rec_sugar.ML |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
2 |
Author: Lorenz Panny, TU Muenchen |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
3 |
Copyright 2013 |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
4 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
5 |
Recursor and corecursor sugar. |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
6 |
*) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
7 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
8 |
signature BNF_FP_REC_SUGAR = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
9 |
sig |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
10 |
val add_primrec_cmd: (binding * string option * mixfix) list -> |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
11 |
(Attrib.binding * string) list -> local_theory -> local_theory; |
53753
ae7f50e70c09
renamed "primcorec" to "primcorecursive", to open the door to a 'theory -> theory' command called "primcorec" (cf. "fun" vs. "function")
blanchet
parents:
53744
diff
changeset
|
12 |
val add_primcorecursive_cmd: bool -> |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
13 |
(binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list -> |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
14 |
Proof.context -> Proof.state |
53822 | 15 |
val add_primcorec_cmd: bool -> |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
16 |
(binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list -> |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
17 |
local_theory -> local_theory |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
18 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
19 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
20 |
structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
21 |
struct |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
22 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
23 |
open BNF_Util |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
24 |
open BNF_FP_Util |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
25 |
open BNF_FP_Rec_Sugar_Util |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
26 |
open BNF_FP_Rec_Sugar_Tactics |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
27 |
|
53811
2967fa35d89e
set code and nitpick_simp attributes on primcorec theorems
blanchet
parents:
53797
diff
changeset
|
28 |
val codeN = "code" |
53791 | 29 |
val ctrN = "ctr" |
30 |
val discN = "disc" |
|
31 |
val selN = "sel" |
|
32 |
||
53811
2967fa35d89e
set code and nitpick_simp attributes on primcorec theorems
blanchet
parents:
53797
diff
changeset
|
33 |
val nitpick_attrs = @{attributes [nitpick_simp]}; |
2967fa35d89e
set code and nitpick_simp attributes on primcorec theorems
blanchet
parents:
53797
diff
changeset
|
34 |
val code_nitpick_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs; |
53794 | 35 |
val simp_attrs = @{attributes [simp]}; |
36 |
||
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
37 |
exception Primrec_Error of string * term list; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
38 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
39 |
fun primrec_error str = raise Primrec_Error (str, []); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
40 |
fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
41 |
fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
42 |
|
53358 | 43 |
fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x); |
44 |
||
53357 | 45 |
val free_name = try (fn Free (v, _) => v); |
46 |
val const_name = try (fn Const (v, _) => v); |
|
53358 | 47 |
val undef_const = Const (@{const_name undefined}, dummyT); |
53357 | 48 |
|
53358 | 49 |
fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1))) |
53720 | 50 |
|> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n); |
53401 | 51 |
val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple; |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
52 |
fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev, |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
53 |
strip_qnt_body @{const_name all} t) |
53720 | 54 |
fun abstract vs = |
55 |
let fun a n (t $ u) = a n t $ a n u |
|
56 |
| a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b) |
|
57 |
| a n t = let val idx = find_index (equal t) vs in |
|
58 |
if idx < 0 then t else Bound (n + idx) end |
|
59 |
in a 0 end; |
|
53735 | 60 |
fun mk_prod1 Ts (t, u) = HOLogic.pair_const (fastype_of1 (Ts, t)) (fastype_of1 (Ts, u)) $ t $ u; |
61 |
fun mk_tuple1 Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 Ts)); |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
62 |
|
53794 | 63 |
fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes |
64 |
|> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE) |
|
65 |
|> map_filter I; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
66 |
|
53310 | 67 |
|
68 |
(* Primrec *) |
|
69 |
||
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
70 |
type eqn_data = { |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
71 |
fun_name: string, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
72 |
rec_type: typ, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
73 |
ctr: term, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
74 |
ctr_args: term list, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
75 |
left_args: term list, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
76 |
right_args: term list, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
77 |
res_type: typ, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
78 |
rhs_term: term, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
79 |
user_eqn: term |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
80 |
}; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
81 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
82 |
fun dissect_eqn lthy fun_names eqn' = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
83 |
let |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
84 |
val eqn = drop_All eqn' |> HOLogic.dest_Trueprop |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
85 |
handle TERM _ => |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
86 |
primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn'; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
87 |
val (lhs, rhs) = HOLogic.dest_eq eqn |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
88 |
handle TERM _ => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
89 |
primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
90 |
val (fun_name, args) = strip_comb lhs |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
91 |
|>> (fn x => if is_Free x then fst (dest_Free x) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
92 |
else primrec_error_eqn "malformed function equation (does not start with free)" eqn); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
93 |
val (left_args, rest) = take_prefix is_Free args; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
94 |
val (nonfrees, right_args) = take_suffix is_Free rest; |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
95 |
val num_nonfrees = length nonfrees; |
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
96 |
val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
97 |
primrec_error_eqn "constructor pattern missing in left-hand side" eqn else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
98 |
primrec_error_eqn "more than one non-variable argument in left-hand side" eqn; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
99 |
val _ = member (op =) fun_names fun_name orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
100 |
primrec_error_eqn "malformed function equation (does not start with function name)" eqn |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
101 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
102 |
val (ctr, ctr_args) = strip_comb (the_single nonfrees); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
103 |
val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
104 |
primrec_error_eqn "partially applied constructor in pattern" eqn; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
105 |
val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
106 |
primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^ |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
107 |
"\" in left-hand side") eqn end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
108 |
val _ = forall is_Free ctr_args orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
109 |
primrec_error_eqn "non-primitive pattern in left-hand side" eqn; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
110 |
val _ = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
111 |
let val b = fold_aterms (fn x as Free (v, _) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
112 |
if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
113 |
not (member (op =) fun_names v) andalso |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
114 |
not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs [] |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
115 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
116 |
null b orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
117 |
primrec_error_eqn ("extra variable(s) in right-hand side: " ^ |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
118 |
commas (map (Syntax.string_of_term lthy) b)) eqn |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
119 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
120 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
121 |
{fun_name = fun_name, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
122 |
rec_type = body_type (type_of ctr), |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
123 |
ctr = ctr, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
124 |
ctr_args = ctr_args, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
125 |
left_args = left_args, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
126 |
right_args = right_args, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
127 |
res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
128 |
rhs_term = rhs, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
129 |
user_eqn = eqn'} |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
130 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
131 |
|
53401 | 132 |
fun rewrite_map_arg get_ctr_pos rec_type res_type = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
133 |
let |
53357 | 134 |
val pT = HOLogic.mk_prodT (rec_type, res_type); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
135 |
|
53357 | 136 |
val maybe_suc = Option.map (fn x => x + 1); |
137 |
fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT) |
|
138 |
| subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b) |
|
139 |
| subst d t = |
|
53358 | 140 |
let |
141 |
val (u, vs) = strip_comb t; |
|
53401 | 142 |
val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1; |
53358 | 143 |
in |
53401 | 144 |
if ctr_pos >= 0 then |
53357 | 145 |
if d = SOME ~1 andalso length vs = ctr_pos then |
146 |
list_comb (permute_args ctr_pos (snd_const pT), vs) |
|
147 |
else if length vs > ctr_pos andalso is_some d |
|
148 |
andalso d = try (fn Bound n => n) (nth vs ctr_pos) then |
|
149 |
list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs)) |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
150 |
else |
53357 | 151 |
primrec_error_eqn ("recursive call not directly applied to constructor argument") t |
152 |
else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then |
|
153 |
list_comb (map_types (K dummyT) u, map2 subst [NONE, d] vs) |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
154 |
else |
53357 | 155 |
list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
156 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
157 |
in |
53357 | 158 |
subst (SOME ~1) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
159 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
160 |
|
53401 | 161 |
fun subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls t = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
162 |
let |
53350 | 163 |
fun subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b) |
53358 | 164 |
| subst bound_Ts (t as g' $ y) = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
165 |
let |
53350 | 166 |
val maybe_direct_y' = AList.lookup (op =) direct_calls y; |
167 |
val maybe_indirect_y' = AList.lookup (op =) indirect_calls y; |
|
53358 | 168 |
val (g, g_args) = strip_comb g'; |
53401 | 169 |
val ctr_pos = try (get_ctr_pos o the) (free_name g) |> the_default ~1; |
170 |
val _ = ctr_pos < 0 orelse length g_args >= ctr_pos orelse |
|
53358 | 171 |
primrec_error_eqn "too few arguments in recursive call" t; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
172 |
in |
53358 | 173 |
if not (member (op =) ctr_args y) then |
174 |
pairself (subst bound_Ts) (g', y) |> (op $) |
|
53401 | 175 |
else if ctr_pos >= 0 then |
53358 | 176 |
list_comb (the maybe_direct_y', g_args) |
53350 | 177 |
else if is_some maybe_indirect_y' then |
53358 | 178 |
(if has_call g' then t else y) |
179 |
|> massage_indirect_rec_call lthy has_call |
|
53401 | 180 |
(rewrite_map_arg get_ctr_pos) bound_Ts y (the maybe_indirect_y') |
53358 | 181 |
|> (if has_call g' then I else curry (op $) g') |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
182 |
else |
53350 | 183 |
t |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
184 |
end |
53350 | 185 |
| subst _ t = t |
186 |
in |
|
187 |
subst [] t |
|
53358 | 188 |
|> tap (fn u => has_call u andalso (* FIXME detect this case earlier *) |
189 |
primrec_error_eqn "recursive call not directly applied to constructor argument" t) |
|
53350 | 190 |
end; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
191 |
|
53358 | 192 |
fun build_rec_arg lthy funs_data has_call ctr_spec maybe_eqn_data = |
193 |
if is_none maybe_eqn_data then undef_const else |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
194 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
195 |
val eqn_data = the maybe_eqn_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
196 |
val t = #rhs_term eqn_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
197 |
val ctr_args = #ctr_args eqn_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
198 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
199 |
val calls = #calls ctr_spec; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
200 |
val n_args = fold (curry (op +) o (fn Direct_Rec _ => 2 | _ => 1)) calls 0; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
201 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
202 |
val no_calls' = tag_list 0 calls |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
203 |
|> map_filter (try (apsnd (fn No_Rec n => n | Direct_Rec (n, _) => n))); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
204 |
val direct_calls' = tag_list 0 calls |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
205 |
|> map_filter (try (apsnd (fn Direct_Rec (_, n) => n))); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
206 |
val indirect_calls' = tag_list 0 calls |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
207 |
|> map_filter (try (apsnd (fn Indirect_Rec n => n))); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
208 |
|
53901 | 209 |
fun make_direct_type _ = dummyT; (* FIXME? *) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
210 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
211 |
val rec_res_type_list = map (fn (x :: _) => (#rec_type x, #res_type x)) funs_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
212 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
213 |
fun make_indirect_type (Type (Tname, Ts)) = Type (Tname, Ts |> map (fn T => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
214 |
let val maybe_res_type = AList.lookup (op =) rec_res_type_list T in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
215 |
if is_some maybe_res_type |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
216 |
then HOLogic.mk_prodT (T, the maybe_res_type) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
217 |
else make_indirect_type T end)) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
218 |
| make_indirect_type T = T; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
219 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
220 |
val args = replicate n_args ("", dummyT) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
221 |
|> Term.rename_wrt_term t |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
222 |
|> map Free |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
223 |
|> fold (fn (ctr_arg_idx, arg_idx) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
224 |
nth_map arg_idx (K (nth ctr_args ctr_arg_idx))) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
225 |
no_calls' |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
226 |
|> fold (fn (ctr_arg_idx, arg_idx) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
227 |
nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_direct_type))) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
228 |
direct_calls' |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
229 |
|> fold (fn (ctr_arg_idx, arg_idx) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
230 |
nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_indirect_type))) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
231 |
indirect_calls'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
232 |
|
53401 | 233 |
val fun_name_ctr_pos_list = |
234 |
map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data; |
|
235 |
val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
236 |
val direct_calls = map (apfst (nth ctr_args) o apsnd (nth args)) direct_calls'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
237 |
val indirect_calls = map (apfst (nth ctr_args) o apsnd (nth args)) indirect_calls'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
238 |
|
53401 | 239 |
val abstractions = args @ #left_args eqn_data @ #right_args eqn_data; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
240 |
in |
53350 | 241 |
t |
53401 | 242 |
|> subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls |
243 |
|> fold_rev lambda abstractions |
|
53350 | 244 |
end; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
245 |
|
53358 | 246 |
fun build_defs lthy bs mxs funs_data rec_specs has_call = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
247 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
248 |
val n_funs = length funs_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
249 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
250 |
val ctr_spec_eqn_data_list' = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
251 |
(take n_funs rec_specs |> map #ctr_specs) ~~ funs_data |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
252 |
|> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y)) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
253 |
##> (fn x => null x orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
254 |
primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
255 |
val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
256 |
primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x)); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
257 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
258 |
val ctr_spec_eqn_data_list = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
259 |
ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair [])); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
260 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
261 |
val recs = take n_funs rec_specs |> map #recx; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
262 |
val rec_args = ctr_spec_eqn_data_list |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
263 |
|> sort ((op <) o pairself (#offset o fst) |> make_ord) |
53358 | 264 |
|> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single)); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
265 |
val ctr_poss = map (fn x => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
266 |
if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
267 |
primrec_error ("inconstant constructor pattern position for function " ^ |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
268 |
quote (#fun_name (hd x))) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
269 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
270 |
hd x |> #left_args |> length) funs_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
271 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
272 |
(recs, ctr_poss) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
273 |
|-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
274 |
|> Syntax.check_terms lthy |
53352 | 275 |
|> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
276 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
277 |
|
53358 | 278 |
fun find_rec_calls has_call eqn_data = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
279 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
280 |
fun find (Abs (_, _, b)) ctr_arg = find b ctr_arg |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
281 |
| find (t as _ $ _) ctr_arg = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
282 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
283 |
val (f', args') = strip_comb t; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
284 |
val n = find_index (equal ctr_arg) args'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
285 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
286 |
if n < 0 then |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
287 |
find f' ctr_arg @ maps (fn x => find x ctr_arg) args' |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
288 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
289 |
let val (f, args) = chop n args' |>> curry list_comb f' in |
53358 | 290 |
if has_call f then |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
291 |
f :: maps (fn x => find x ctr_arg) args |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
292 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
293 |
find f ctr_arg @ maps (fn x => find x ctr_arg) args |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
294 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
295 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
296 |
| find _ _ = []; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
297 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
298 |
map (find (#rhs_term eqn_data)) (#ctr_args eqn_data) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
299 |
|> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss)) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
300 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
301 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
302 |
fun add_primrec fixes specs lthy = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
303 |
let |
53352 | 304 |
val (bs, mxs) = map_split (apfst fst) fixes; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
305 |
val fun_names = map Binding.name_of bs; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
306 |
val eqns_data = map (snd #> dissect_eqn lthy fun_names) specs; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
307 |
val funs_data = eqns_data |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
308 |
|> partition_eq ((op =) o pairself #fun_name) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
309 |
|> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
310 |
|> map (fn (x, y) => the_single y handle List.Empty => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
311 |
primrec_error ("missing equations for function " ^ quote x)); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
312 |
|
53358 | 313 |
val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =)); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
314 |
val arg_Ts = map (#rec_type o hd) funs_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
315 |
val res_Ts = map (#res_type o hd) funs_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
316 |
val callssss = funs_data |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
317 |
|> map (partition_eq ((op =) o pairself #ctr)) |
53358 | 318 |
|> map (maps (map_filter (find_rec_calls has_call))); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
319 |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
320 |
val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') = |
53794 | 321 |
rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
322 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
323 |
val actual_nn = length funs_data; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
324 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
325 |
val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
326 |
map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
327 |
primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^ |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
328 |
" is not a constructor in left-hand side") user_eqn) eqns_data end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
329 |
|
53358 | 330 |
val defs = build_defs lthy' bs mxs funs_data rec_specs has_call; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
331 |
|
53329 | 332 |
fun prove def_thms' {ctr_specs, nested_map_idents, nested_map_comps, ...} induct_thm fun_data |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
333 |
lthy = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
334 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
335 |
val fun_name = #fun_name (hd fun_data); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
336 |
val def_thms = map (snd o snd) def_thms'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
337 |
val simp_thms = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
338 |
|> fst |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
339 |
|> map_filter (try (fn (x, [y]) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
340 |
(#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y))) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
341 |
|> map (fn (user_eqn, num_extra_args, rec_thm) => |
53329 | 342 |
mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
343 |
|> K |> Goal.prove lthy [] [] user_eqn) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
344 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
345 |
val notes = |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
346 |
[(inductN, if n2m then [induct_thm] else [], []), |
53811
2967fa35d89e
set code and nitpick_simp attributes on primcorec theorems
blanchet
parents:
53797
diff
changeset
|
347 |
(simpsN, simp_thms, code_nitpick_simp_attrs @ simp_attrs)] |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
348 |
|> filter_out (null o #2) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
349 |
|> map (fn (thmN, thms, attrs) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
350 |
((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])])); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
351 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
352 |
lthy |> Local_Theory.notes notes |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
353 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
354 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
355 |
val common_name = mk_common_name fun_names; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
356 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
357 |
val common_notes = |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
358 |
[(inductN, if n2m then [induct_thm] else [], [])] |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
359 |
|> filter_out (null o #2) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
360 |
|> map (fn (thmN, thms, attrs) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
361 |
((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])])); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
362 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
363 |
lthy' |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
364 |
|> fold_map Local_Theory.define defs |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
365 |
|-> snd oo (fn def_thms' => fold_map3 (prove def_thms') (take actual_nn rec_specs) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
366 |
(take actual_nn induct_thms) funs_data) |
53797 | 367 |
|> Local_Theory.notes common_notes |> snd |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
368 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
369 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
370 |
fun add_primrec_cmd raw_fixes raw_specs lthy = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
371 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
372 |
val _ = let val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes) in null d orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
373 |
primrec_error ("duplicate function name(s): " ^ commas d) end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
374 |
val (fixes, specs) = fst (Specification.read_spec raw_fixes raw_specs lthy); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
375 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
376 |
add_primrec fixes specs lthy |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
377 |
handle ERROR str => primrec_error str |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
378 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
379 |
handle Primrec_Error (str, eqns) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
380 |
if null eqns |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
381 |
then error ("primrec_new error:\n " ^ str) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
382 |
else error ("primrec_new error:\n " ^ str ^ "\nin\n " ^ |
53822 | 383 |
space_implode "\n " (map (quote o Syntax.string_of_term lthy) eqns)); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
384 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
385 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
386 |
|
53310 | 387 |
(* Primcorec *) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
388 |
|
53341 | 389 |
type co_eqn_data_disc = { |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
390 |
fun_name: string, |
53720 | 391 |
fun_T: typ, |
53401 | 392 |
fun_args: term list, |
53720 | 393 |
ctr: term, |
53341 | 394 |
ctr_no: int, (*###*) |
53720 | 395 |
disc: term, |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
396 |
prems: term list, |
53822 | 397 |
auto_gen: bool, |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
398 |
user_eqn: term |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
399 |
}; |
53341 | 400 |
type co_eqn_data_sel = { |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
401 |
fun_name: string, |
53720 | 402 |
fun_T: typ, |
53401 | 403 |
fun_args: term list, |
53341 | 404 |
ctr: term, |
405 |
sel: term, |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
406 |
rhs_term: term, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
407 |
user_eqn: term |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
408 |
}; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
409 |
datatype co_eqn_data = |
53341 | 410 |
Disc of co_eqn_data_disc | |
411 |
Sel of co_eqn_data_sel; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
412 |
|
53720 | 413 |
fun co_dissect_eqn_disc sequential fun_names corec_specs prems' concl matchedsss = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
414 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
415 |
fun find_subterm p = let (* FIXME \<exists>? *) |
53401 | 416 |
fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
417 |
| f t = if p t then SOME t else NONE |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
418 |
in f end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
419 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
420 |
val applied_fun = concl |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
421 |
|> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of)) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
422 |
|> the |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
423 |
handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl; |
53720 | 424 |
val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free; |
425 |
val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name); |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
426 |
|
53720 | 427 |
val discs = map #disc ctr_specs; |
428 |
val ctrs = map #ctr ctr_specs; |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
429 |
val not_disc = head_of concl = @{term Not}; |
53401 | 430 |
val _ = not_disc andalso length ctrs <> 2 andalso |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
431 |
primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
432 |
val disc = find_subterm (member (op =) discs o head_of) concl; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
433 |
val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
434 |
|> (fn SOME t => let val n = find_index (equal t) ctrs in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
435 |
if n >= 0 then SOME n else NONE end | _ => NONE); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
436 |
val _ = is_some disc orelse is_some eq_ctr0 orelse |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
437 |
primrec_error_eqn "no discriminator in equation" concl; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
438 |
val ctr_no' = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
439 |
if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
440 |
val ctr_no = if not_disc then 1 - ctr_no' else ctr_no'; |
53720 | 441 |
val ctr = #ctr (nth ctr_specs ctr_no); |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
442 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
443 |
val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_; |
53720 | 444 |
val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default []; |
445 |
val prems = map (abstract (List.rev fun_args)) prems'; |
|
446 |
val real_prems = |
|
53878 | 447 |
(if catch_all orelse sequential then maps negate_disj matchedss else []) @ |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
448 |
(if catch_all then [] else prems); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
449 |
|
53720 | 450 |
val matchedsss' = AList.delete (op =) fun_name matchedsss |
451 |
|> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [real_prems]); |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
452 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
453 |
val user_eqn = |
53720 | 454 |
(real_prems, betapply (#disc (nth ctr_specs ctr_no), applied_fun)) |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
455 |
|>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
456 |
|> Logic.list_implies; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
457 |
in |
53341 | 458 |
(Disc { |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
459 |
fun_name = fun_name, |
53720 | 460 |
fun_T = fun_T, |
53401 | 461 |
fun_args = fun_args, |
53720 | 462 |
ctr = ctr, |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
463 |
ctr_no = ctr_no, |
53720 | 464 |
disc = #disc (nth ctr_specs ctr_no), |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
465 |
prems = real_prems, |
53822 | 466 |
auto_gen = catch_all, |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
467 |
user_eqn = user_eqn |
53720 | 468 |
}, matchedsss') |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
469 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
470 |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
471 |
fun co_dissect_eqn_sel fun_names corec_specs eqn' of_spec eqn = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
472 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
473 |
val (lhs, rhs) = HOLogic.dest_eq eqn |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
474 |
handle TERM _ => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
475 |
primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
476 |
val sel = head_of lhs; |
53720 | 477 |
val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
478 |
handle TERM _ => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
479 |
primrec_error_eqn "malformed selector argument in left-hand side" eqn; |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
480 |
val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
481 |
handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn; |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
482 |
val ctr_spec = |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
483 |
if is_some of_spec |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
484 |
then the (find_first (equal (the of_spec) o #ctr) (#ctr_specs corec_spec)) |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
485 |
else #ctr_specs corec_spec |> filter (exists (equal sel) o #sels) |> the_single |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
486 |
handle List.Empty => primrec_error_eqn "ambiguous selector - use \"of\"" eqn; |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
487 |
val user_eqn = drop_All eqn'; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
488 |
in |
53341 | 489 |
Sel { |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
490 |
fun_name = fun_name, |
53720 | 491 |
fun_T = fun_T, |
53401 | 492 |
fun_args = fun_args, |
53341 | 493 |
ctr = #ctr ctr_spec, |
494 |
sel = sel, |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
495 |
rhs_term = rhs, |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
496 |
user_eqn = user_eqn |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
497 |
} |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
498 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
499 |
|
53720 | 500 |
fun co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss = |
53910 | 501 |
let |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
502 |
val (lhs, rhs) = HOLogic.dest_eq imp_rhs; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
503 |
val fun_name = head_of lhs |> fst o dest_Free; |
53720 | 504 |
val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
505 |
val (ctr, ctr_args) = strip_comb rhs; |
53720 | 506 |
val {disc, sels, ...} = the (find_first (equal ctr o #ctr) ctr_specs) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
507 |
handle Option.Option => primrec_error_eqn "not a constructor" ctr; |
53341 | 508 |
|
53720 | 509 |
val disc_imp_rhs = betapply (disc, lhs); |
510 |
val (maybe_eqn_data_disc, matchedsss') = if length ctr_specs = 1 |
|
511 |
then (NONE, matchedsss) |
|
53341 | 512 |
else apfst SOME (co_dissect_eqn_disc |
53720 | 513 |
sequential fun_names corec_specs imp_prems disc_imp_rhs matchedsss); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
514 |
|
53720 | 515 |
val sel_imp_rhss = (sels ~~ ctr_args) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
516 |
|> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg)); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
517 |
|
53856 | 518 |
(* |
53360 | 519 |
val _ = tracing ("reduced\n " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n \<cdot> " ^ |
53341 | 520 |
(is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n \<cdot> ")) "" ^ |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
521 |
space_implode "\n \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss)); |
53856 | 522 |
*) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
523 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
524 |
val eqns_data_sel = |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
525 |
map (co_dissect_eqn_sel fun_names corec_specs eqn' (SOME ctr)) sel_imp_rhss; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
526 |
in |
53720 | 527 |
(the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss') |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
528 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
529 |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
530 |
fun co_dissect_eqn sequential fun_names corec_specs eqn' of_spec matchedsss = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
531 |
let |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
532 |
val eqn = drop_All eqn' |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
533 |
handle TERM _ => primrec_error_eqn "malformed function equation" eqn'; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
534 |
val (imp_prems, imp_rhs) = Logic.strip_horn eqn |
53341 | 535 |
|> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
536 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
537 |
val head = imp_rhs |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
538 |
|> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq)) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
539 |
|> head_of; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
540 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
541 |
val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
542 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
543 |
val discs = maps #ctr_specs corec_specs |> map #disc; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
544 |
val sels = maps #ctr_specs corec_specs |> maps #sels; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
545 |
val ctrs = maps #ctr_specs corec_specs |> map #ctr; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
546 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
547 |
if member (op =) discs head orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
548 |
is_some maybe_rhs andalso |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
549 |
member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then |
53720 | 550 |
co_dissect_eqn_disc sequential fun_names corec_specs imp_prems imp_rhs matchedsss |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
551 |
|>> single |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
552 |
else if member (op =) sels head then |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
553 |
([co_dissect_eqn_sel fun_names corec_specs eqn' of_spec imp_rhs], matchedsss) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
554 |
else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then |
53720 | 555 |
co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
556 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
557 |
primrec_error_eqn "malformed function equation" eqn |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
558 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
559 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
560 |
fun build_corec_arg_disc ctr_specs {fun_args, ctr_no, prems, ...} = |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
561 |
if is_none (#pred (nth ctr_specs ctr_no)) then I else |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
562 |
mk_conjs prems |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
563 |
|> curry subst_bounds (List.rev fun_args) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
564 |
|> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
565 |
|> K |> nth_map (the (#pred (nth ctr_specs ctr_no))); |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
566 |
|
53720 | 567 |
fun build_corec_arg_no_call sel_eqns sel = |
568 |
find_first (equal sel o #sel) sel_eqns |
|
569 |
|> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term) |
|
570 |
|> the_default undef_const |
|
53411 | 571 |
|> K; |
53360 | 572 |
|
53735 | 573 |
fun build_corec_args_direct_call lthy has_call sel_eqns sel = |
53360 | 574 |
let |
53411 | 575 |
val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns; |
53360 | 576 |
in |
53876 | 577 |
if is_none maybe_sel_eqn then (I, I, I) else |
578 |
let |
|
579 |
val {fun_args, rhs_term, ... } = the maybe_sel_eqn; |
|
53890 | 580 |
fun rewrite_q _ t = if has_call t then @{term False} else @{term True}; |
581 |
fun rewrite_g _ t = if has_call t then undef_const else t; |
|
53899 | 582 |
fun rewrite_h bound_Ts t = |
583 |
if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const; |
|
53889 | 584 |
fun massage f t = massage_direct_corec_call lthy has_call f [] rhs_term |> abs_tuple fun_args; |
53876 | 585 |
in |
586 |
(massage rewrite_q, |
|
587 |
massage rewrite_g, |
|
588 |
massage rewrite_h) |
|
589 |
end |
|
53360 | 590 |
end; |
591 |
||
53411 | 592 |
fun build_corec_arg_indirect_call lthy has_call sel_eqns sel = |
593 |
let |
|
594 |
val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns; |
|
53899 | 595 |
in |
596 |
if is_none maybe_sel_eqn then I else |
|
597 |
let |
|
598 |
val {fun_args, rhs_term, ...} = the maybe_sel_eqn; |
|
599 |
fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b) |
|
600 |
| rewrite bound_Ts U T (t as _ $ _) = |
|
601 |
let val (u, vs) = strip_comb t in |
|
602 |
if is_Free u andalso has_call u then |
|
603 |
Inr_const U T $ mk_tuple1 bound_Ts vs |
|
604 |
else if try (fst o dest_Const) u = SOME @{const_name prod_case} then |
|
605 |
map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb |
|
606 |
else |
|
607 |
list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs) |
|
608 |
end |
|
609 |
| rewrite _ U T t = |
|
610 |
if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t; |
|
611 |
fun massage t = |
|
53890 | 612 |
massage_indirect_corec_call lthy has_call rewrite [] (range_type (fastype_of t)) rhs_term |
53735 | 613 |
|> abs_tuple fun_args; |
53899 | 614 |
in |
615 |
massage |
|
616 |
end |
|
53411 | 617 |
end; |
53360 | 618 |
|
619 |
fun build_corec_args_sel lthy has_call all_sel_eqns ctr_spec = |
|
53341 | 620 |
let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in |
621 |
if null sel_eqns then I else |
|
622 |
let |
|
623 |
val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec; |
|
624 |
||
625 |
val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list; |
|
626 |
val direct_calls' = map_filter (try (apsnd (fn Direct_Corec n => n))) sel_call_list; |
|
627 |
val indirect_calls' = map_filter (try (apsnd (fn Indirect_Corec n => n))) sel_call_list; |
|
628 |
in |
|
53360 | 629 |
I |
53735 | 630 |
#> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls' |
53360 | 631 |
#> fold (fn (sel, (q, g, h)) => |
53735 | 632 |
let val (fq, fg, fh) = build_corec_args_direct_call lthy has_call sel_eqns sel in |
633 |
nth_map q fq o nth_map g fg o nth_map h fh end) direct_calls' |
|
53360 | 634 |
#> fold (fn (sel, n) => nth_map n |
53411 | 635 |
(build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls' |
53341 | 636 |
end |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
637 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
638 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
639 |
fun co_build_defs lthy bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss = |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
640 |
let |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
641 |
val corec_specs' = take (length bs) corec_specs; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
642 |
val corecs = map #corec corec_specs'; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
643 |
val ctr_specss = map #ctr_specs corec_specs'; |
53360 | 644 |
val corec_args = hd corecs |
645 |
|> fst o split_last o binder_types o fastype_of |
|
646 |
|> map (Const o pair @{const_name undefined}) |
|
53720 | 647 |
|> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss |
53360 | 648 |
|> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss; |
53735 | 649 |
fun currys [] t = t |
650 |
| currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0)) |
|
651 |
|> fold_rev (Term.abs o pair Name.uu) Ts; |
|
53401 | 652 |
|
53856 | 653 |
(* |
53360 | 654 |
val _ = tracing ("corecursor arguments:\n \<cdot> " ^ |
53411 | 655 |
space_implode "\n \<cdot> " (map (Syntax.string_of_term lthy) corec_args)); |
53856 | 656 |
*) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
657 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
658 |
val exclss' = |
53720 | 659 |
disc_eqnss |
53822 | 660 |
|> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x)) |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
661 |
#> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs []) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
662 |
#> maps (uncurry (map o pair) |
53822 | 663 |
#> map (fn ((fun_args, c, x, a), (_, c', y, a')) => |
664 |
((c, c', a orelse a'), (x, s_not (mk_conjs y))) |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
665 |
||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
666 |
||> Logic.list_implies |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
667 |
||> curry Logic.list_all (map dest_Free fun_args)))) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
668 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
669 |
map (list_comb o rpair corec_args) corecs |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
670 |
|> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
671 |
|> map2 currys arg_Tss |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
672 |
|> Syntax.check_terms lthy |
53352 | 673 |
|> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
674 |
|> rpair exclss' |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
675 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
676 |
|
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
677 |
fun mk_real_disc_eqns fun_binding arg_Ts {ctr_specs, ...} sel_eqns disc_eqns = |
53720 | 678 |
if length disc_eqns <> length ctr_specs - 1 then disc_eqns else |
679 |
let |
|
680 |
val n = 0 upto length ctr_specs |
|
681 |
|> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns)); |
|
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
682 |
val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns) |
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
683 |
|> the_default (map (curry Free Name.uu) arg_Ts) o merge_options; |
53720 | 684 |
val extra_disc_eqn = { |
685 |
fun_name = Binding.name_of fun_binding, |
|
686 |
fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))), |
|
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
687 |
fun_args = fun_args, |
53720 | 688 |
ctr = #ctr (nth ctr_specs n), |
689 |
ctr_no = n, |
|
690 |
disc = #disc (nth ctr_specs n), |
|
53878 | 691 |
prems = maps (negate_conj o #prems) disc_eqns, |
53822 | 692 |
auto_gen = true, |
53720 | 693 |
user_eqn = undef_const}; |
694 |
in |
|
695 |
chop n disc_eqns ||> cons extra_disc_eqn |> (op @) |
|
696 |
end; |
|
697 |
||
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
698 |
fun add_primcorec simple sequential fixes specs of_specs lthy = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
699 |
let |
53352 | 700 |
val (bs, mxs) = map_split (apfst fst) fixes; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
701 |
val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
702 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
703 |
val callssss = []; (* FIXME *) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
704 |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
705 |
val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms, |
53797 | 706 |
strong_coinduct_thms), lthy') = |
53794 | 707 |
corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
708 |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
709 |
val actual_nn = length bs; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
710 |
val fun_names = map Binding.name_of bs; |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
711 |
val corec_specs = take actual_nn corec_specs'; (*###*) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
712 |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
713 |
val eqns_data = |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
714 |
fold_map2 (co_dissect_eqn sequential fun_names corec_specs) (map snd specs) of_specs [] |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
715 |
|> flat o fst; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
716 |
|
53720 | 717 |
val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
718 |
|> partition_eq ((op =) o pairself #fun_name) |
53720 | 719 |
|> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
720 |
|> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd); |
53720 | 721 |
val _ = disc_eqnss' |> map (fn x => |
722 |
let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse |
|
723 |
primrec_error_eqns "excess discriminator equations in definition" |
|
724 |
(maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end); |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
725 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
726 |
val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
727 |
|> partition_eq ((op =) o pairself #fun_name) |
53720 | 728 |
|> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
729 |
|> map (flat o snd); |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
730 |
|
53360 | 731 |
val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =)); |
732 |
val arg_Tss = map (binder_types o snd o fst) fixes; |
|
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
733 |
val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss'; |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
734 |
val (defs, exclss') = |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
735 |
co_build_defs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
736 |
|
53822 | 737 |
fun prove_excl_tac (c, c', a) = |
53903 | 738 |
if a orelse c = c' orelse sequential then SOME (K (mk_primcorec_assumption_tac lthy [])) |
53822 | 739 |
else if simple then SOME (K (auto_tac lthy)) |
740 |
else NONE; |
|
741 |
||
53856 | 742 |
(* |
53822 | 743 |
val _ = tracing ("exclusiveness properties:\n \<cdot> " ^ |
744 |
space_implode "\n \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss')); |
|
53856 | 745 |
*) |
53822 | 746 |
|
747 |
val exclss'' = exclss' |> map (map (fn (idx, t) => |
|
748 |
(idx, (Option.map (Goal.prove lthy [] [] t) (prove_excl_tac idx), t)))); |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
749 |
val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss''; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
750 |
val (obligation_idxss, obligationss) = exclss'' |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
751 |
|> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd)) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
752 |
|> split_list o map split_list; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
753 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
754 |
fun prove thmss' def_thms' lthy = |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
755 |
let |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
756 |
val def_thms = map (snd o snd) def_thms'; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
757 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
758 |
val exclss' = map (op ~~) (obligation_idxss ~~ thmss'); |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
759 |
fun mk_exclsss excls n = |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
760 |
(excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1)) |
53822 | 761 |
|-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm]))); |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
762 |
val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
763 |
|-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs)); |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
764 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
765 |
fun prove_disc {ctr_specs, ...} exclsss |
53720 | 766 |
{fun_name, fun_T, fun_args, ctr_no, prems, user_eqn, ...} = |
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
767 |
if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else |
53720 | 768 |
let |
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
769 |
val {disc_corec, ...} = nth ctr_specs ctr_no; |
53720 | 770 |
val k = 1 + ctr_no; |
771 |
val m = length prems; |
|
772 |
val t = |
|
773 |
list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0)) |
|
774 |
|> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*) |
|
775 |
|> HOLogic.mk_Trueprop |
|
776 |
|> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) |
|
777 |
|> curry Logic.list_all (map dest_Free fun_args); |
|
778 |
in |
|
779 |
mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss |
|
780 |
|> K |> Goal.prove lthy [] [] t |
|
781 |
|> pair (#disc (nth ctr_specs ctr_no)) |
|
782 |
|> single |
|
783 |
end; |
|
784 |
||
785 |
fun prove_sel {ctr_specs, nested_maps, nested_map_idents, nested_map_comps, ...} |
|
786 |
disc_eqns exclsss {fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} = |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
787 |
let |
53909 | 788 |
val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs; |
53720 | 789 |
val ctr_no = find_index (equal ctr o #ctr) ctr_specs; |
53878 | 790 |
val prems = the_default (maps (negate_conj o #prems) disc_eqns) |
53720 | 791 |
(find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems); |
792 |
val sel_corec = find_index (equal sel) (#sels ctr_spec) |
|
793 |
|> nth (#sel_corecs ctr_spec); |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
794 |
val k = 1 + ctr_no; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
795 |
val m = length prems; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
796 |
val t = |
53720 | 797 |
list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0)) |
798 |
|> curry betapply sel |
|
799 |
|> rpair (abstract (List.rev fun_args) rhs_term) |
|
800 |
|> HOLogic.mk_Trueprop o HOLogic.mk_eq |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
801 |
|> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) |
53720 | 802 |
|> curry Logic.list_all (map dest_Free fun_args); |
53910 | 803 |
val (distincts, _, _, splits, split_asms) = case_thms_of_term lthy [] rhs_term; |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
804 |
in |
53910 | 805 |
mk_primcorec_sel_tac lthy def_thms distincts splits split_asms nested_maps |
806 |
nested_map_idents nested_map_comps sel_corec k m exclsss |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
807 |
|> K |> Goal.prove lthy [] [] t |
53720 | 808 |
|> pair sel |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
809 |
end; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
810 |
|
53791 | 811 |
fun prove_ctr disc_alist sel_alist disc_eqns sel_eqns {ctr, disc, sels, collapse, ...} = |
53720 | 812 |
if not (exists (equal ctr o #ctr) disc_eqns) |
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
813 |
andalso not (exists (equal ctr o #ctr) sel_eqns) |
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
814 |
orelse (* don't try to prove theorems when some sel_eqns are missing *) |
53720 | 815 |
filter (equal ctr o #ctr) sel_eqns |
816 |
|> fst o finds ((op =) o apsnd #sel) sels |
|
817 |
|> exists (null o snd) |
|
818 |
then [] else |
|
819 |
let |
|
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
820 |
val (fun_name, fun_T, fun_args, prems) = |
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
821 |
(find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns) |
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
822 |
|>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x)) |
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
823 |
||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [])) |
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
824 |
|> the o merge_options; |
53720 | 825 |
val m = length prems; |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
826 |
val t = filter (equal ctr o #ctr) sel_eqns |
53720 | 827 |
|> fst o finds ((op =) o apsnd #sel) sels |
828 |
|> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract) |
|
829 |
|> curry list_comb ctr |
|
830 |
|> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T), |
|
831 |
map Bound (length fun_args - 1 downto 0))) |
|
832 |
|> HOLogic.mk_Trueprop |
|
833 |
|> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) |
|
834 |
|> curry Logic.list_all (map dest_Free fun_args); |
|
53791 | 835 |
val maybe_disc_thm = AList.lookup (op =) disc_alist disc; |
836 |
val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist); |
|
53720 | 837 |
in |
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
838 |
mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms |
53720 | 839 |
|> K |> Goal.prove lthy [] [] t |
840 |
|> single |
|
53876 | 841 |
end; |
53720 | 842 |
|
53791 | 843 |
val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss; |
844 |
val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss; |
|
53744
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
845 |
|
53791 | 846 |
val disc_thmss = map (map snd) disc_alists; |
847 |
val sel_thmss = map (map snd) sel_alists; |
|
848 |
val ctr_thmss = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss |
|
53744
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
849 |
(map #ctr_specs corec_specs); |
53791 | 850 |
|
53875
68786e8d1fe6
add non-corecursive constructor view theorems to simps
panny
parents:
53856
diff
changeset
|
851 |
val safess = map (map (not o exists_subterm (member (op =) (map SOME fun_names) o |
68786e8d1fe6
add non-corecursive constructor view theorems to simps
panny
parents:
53856
diff
changeset
|
852 |
try (fst o dest_Free)) o snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o |
68786e8d1fe6
add non-corecursive constructor view theorems to simps
panny
parents:
53856
diff
changeset
|
853 |
Logic.strip_imp_concl o drop_All o prop_of)) ctr_thmss; |
53822 | 854 |
val safe_ctr_thmss = map (map snd o filter fst o (op ~~)) (safess ~~ ctr_thmss); |
53744
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
855 |
|
53875
68786e8d1fe6
add non-corecursive constructor view theorems to simps
panny
parents:
53856
diff
changeset
|
856 |
val simp_thmss = |
68786e8d1fe6
add non-corecursive constructor view theorems to simps
panny
parents:
53856
diff
changeset
|
857 |
map3 (fn x => fn y => fn z => x @ y @ z) disc_thmss sel_thmss safe_ctr_thmss; |
53795 | 858 |
|
53797 | 859 |
val common_name = mk_common_name fun_names; |
860 |
||
53744
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
861 |
val anonymous_notes = |
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
862 |
[(flat safe_ctr_thmss, simp_attrs)] |
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
863 |
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); |
53791 | 864 |
|
865 |
val notes = |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
866 |
[(coinductN, map (if n2m then single else K []) coinduct_thms, []), |
53811
2967fa35d89e
set code and nitpick_simp attributes on primcorec theorems
blanchet
parents:
53797
diff
changeset
|
867 |
(codeN, ctr_thmss(*FIXME*), code_nitpick_simp_attrs), |
53797 | 868 |
(ctrN, ctr_thmss, []), |
53791 | 869 |
(discN, disc_thmss, simp_attrs), |
53795 | 870 |
(selN, sel_thmss, simp_attrs), |
53797 | 871 |
(simpsN, simp_thmss, []), |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
872 |
(strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])] |
53791 | 873 |
|> maps (fn (thmN, thmss, attrs) => |
874 |
map2 (fn fun_name => fn thms => |
|
875 |
((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])])) |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
876 |
fun_names (take actual_nn thmss)) |
53791 | 877 |
|> filter_out (null o fst o hd o snd); |
53797 | 878 |
|
879 |
val common_notes = |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
880 |
[(coinductN, if n2m then [coinduct_thm] else [], []), |
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
881 |
(strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])] |
53797 | 882 |
|> filter_out (null o #2) |
883 |
|> map (fn (thmN, thms, attrs) => |
|
884 |
((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])])); |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
885 |
in |
53797 | 886 |
lthy |> Local_Theory.notes (anonymous_notes @ notes @ common_notes) |> snd |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
887 |
end; |
53822 | 888 |
|
889 |
fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss'; |
|
890 |
||
891 |
val _ = if not simple orelse forall null obligationss then () else |
|
892 |
primrec_error "need exclusiveness proofs - use primcorecursive instead of primcorec"; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
893 |
in |
53822 | 894 |
if simple then |
895 |
lthy' |
|
896 |
|> after_qed (map (fn [] => []) obligationss) |
|
897 |
|> pair NONE o SOME |
|
898 |
else |
|
899 |
lthy' |
|
900 |
|> Proof.theorem NONE after_qed obligationss |
|
901 |
|> Proof.refine (Method.primitive_text I) |
|
902 |
|> Seq.hd |
|
903 |
|> rpair NONE o SOME |
|
904 |
end; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
905 |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
906 |
fun add_primcorec_ursive_cmd simple seq (raw_fixes, raw_specs') lthy = |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
907 |
let |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
908 |
val (raw_specs, of_specs) = split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy)); |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
909 |
val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
910 |
in |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
911 |
add_primcorec simple seq fixes specs of_specs lthy |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
912 |
handle ERROR str => primrec_error str |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
913 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
914 |
handle Primrec_Error (str, eqns) => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
915 |
if null eqns |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
916 |
then error ("primcorec error:\n " ^ str) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
917 |
else error ("primcorec error:\n " ^ str ^ "\nin\n " ^ |
53822 | 918 |
space_implode "\n " (map (quote o Syntax.string_of_term lthy) eqns)); |
919 |
||
920 |
val add_primcorecursive_cmd = (the o fst) ooo add_primcorec_ursive_cmd false; |
|
921 |
val add_primcorec_cmd = (the o snd) ooo add_primcorec_ursive_cmd true; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
922 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
923 |
end; |