author | panny |
Tue, 24 Sep 2013 18:07:09 +0200 | |
changeset 53831 | 80423b9080cf |
parent 53830 | ed2eb7df2aac |
child 53835 | 687116951569 |
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) |
53793
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
54 |
fun s_not @{const True} = @{const False} |
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
55 |
| s_not @{const False} = @{const True} |
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
56 |
| s_not (@{const Not} $ t) = t |
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
57 |
| s_not (@{const Trueprop} $ t) = @{const Trueprop} $ s_not t |
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
58 |
| s_not t = HOLogic.mk_not t |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
59 |
val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True}; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
60 |
val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False}; |
53793
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
61 |
fun invert_prems [t] = map s_not (HOLogic.disjuncts t) |
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
62 |
| invert_prems ts = [mk_disjs (map s_not ts)]; |
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
63 |
fun invert_prems_disj [t] = map s_not (HOLogic.disjuncts t) |
d2f8bca22a51
avoid giving same name to simplifying constructor as to real one (to avoid risks of confusion when reading the code)
blanchet
parents:
53791
diff
changeset
|
64 |
| invert_prems_disj ts = [mk_disjs (map (mk_conjs o map s_not o HOLogic.disjuncts) ts)]; |
53720 | 65 |
fun abstract vs = |
66 |
let fun a n (t $ u) = a n t $ a n u |
|
67 |
| a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b) |
|
68 |
| a n t = let val idx = find_index (equal t) vs in |
|
69 |
if idx < 0 then t else Bound (n + idx) end |
|
70 |
in a 0 end; |
|
53735 | 71 |
fun mk_prod1 Ts (t, u) = HOLogic.pair_const (fastype_of1 (Ts, t)) (fastype_of1 (Ts, u)) $ t $ u; |
72 |
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
|
73 |
|
53794 | 74 |
fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes |
75 |
|> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE) |
|
76 |
|> map_filter I; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
77 |
|
53310 | 78 |
|
79 |
(* Primrec *) |
|
80 |
||
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
81 |
type eqn_data = { |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
82 |
fun_name: string, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
83 |
rec_type: typ, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
84 |
ctr: term, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
85 |
ctr_args: term list, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
86 |
left_args: term list, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
87 |
right_args: term list, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
88 |
res_type: typ, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
89 |
rhs_term: term, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
90 |
user_eqn: term |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
91 |
}; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
92 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
93 |
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
|
94 |
let |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
95 |
val eqn = drop_All eqn' |> HOLogic.dest_Trueprop |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
96 |
handle TERM _ => |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
97 |
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
|
98 |
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
|
99 |
handle TERM _ => |
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 (expected \"lhs = rhs\")" eqn'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
101 |
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
|
102 |
|>> (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
|
103 |
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
|
104 |
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
|
105 |
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
|
106 |
val num_nonfrees = length nonfrees; |
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
107 |
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
|
108 |
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
|
109 |
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
|
110 |
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
|
111 |
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
|
112 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
113 |
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
|
114 |
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
|
115 |
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
|
116 |
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
|
117 |
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
|
118 |
"\" 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
|
119 |
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
|
120 |
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
|
121 |
val _ = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
122 |
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
|
123 |
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
|
124 |
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
|
125 |
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
|
126 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
127 |
null b orelse |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
128 |
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
|
129 |
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
|
130 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
131 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
132 |
{fun_name = fun_name, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
133 |
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
|
134 |
ctr = ctr, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
135 |
ctr_args = ctr_args, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
136 |
left_args = left_args, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
137 |
right_args = right_args, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
138 |
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
|
139 |
rhs_term = rhs, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
140 |
user_eqn = eqn'} |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
141 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
142 |
|
53401 | 143 |
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
|
144 |
let |
53357 | 145 |
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
|
146 |
|
53357 | 147 |
val maybe_suc = Option.map (fn x => x + 1); |
148 |
fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT) |
|
149 |
| subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b) |
|
150 |
| subst d t = |
|
53358 | 151 |
let |
152 |
val (u, vs) = strip_comb t; |
|
53401 | 153 |
val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1; |
53358 | 154 |
in |
53401 | 155 |
if ctr_pos >= 0 then |
53357 | 156 |
if d = SOME ~1 andalso length vs = ctr_pos then |
157 |
list_comb (permute_args ctr_pos (snd_const pT), vs) |
|
158 |
else if length vs > ctr_pos andalso is_some d |
|
159 |
andalso d = try (fn Bound n => n) (nth vs ctr_pos) then |
|
160 |
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
|
161 |
else |
53357 | 162 |
primrec_error_eqn ("recursive call not directly applied to constructor argument") t |
163 |
else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then |
|
164 |
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
|
165 |
else |
53357 | 166 |
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
|
167 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
168 |
in |
53357 | 169 |
subst (SOME ~1) |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
170 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
171 |
|
53401 | 172 |
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
|
173 |
let |
53350 | 174 |
fun subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b) |
53358 | 175 |
| 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
|
176 |
let |
53350 | 177 |
val maybe_direct_y' = AList.lookup (op =) direct_calls y; |
178 |
val maybe_indirect_y' = AList.lookup (op =) indirect_calls y; |
|
53358 | 179 |
val (g, g_args) = strip_comb g'; |
53401 | 180 |
val ctr_pos = try (get_ctr_pos o the) (free_name g) |> the_default ~1; |
181 |
val _ = ctr_pos < 0 orelse length g_args >= ctr_pos orelse |
|
53358 | 182 |
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
|
183 |
in |
53358 | 184 |
if not (member (op =) ctr_args y) then |
185 |
pairself (subst bound_Ts) (g', y) |> (op $) |
|
53401 | 186 |
else if ctr_pos >= 0 then |
53358 | 187 |
list_comb (the maybe_direct_y', g_args) |
53350 | 188 |
else if is_some maybe_indirect_y' then |
53358 | 189 |
(if has_call g' then t else y) |
190 |
|> massage_indirect_rec_call lthy has_call |
|
53401 | 191 |
(rewrite_map_arg get_ctr_pos) bound_Ts y (the maybe_indirect_y') |
53358 | 192 |
|> (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
|
193 |
else |
53350 | 194 |
t |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
195 |
end |
53350 | 196 |
| subst _ t = t |
197 |
in |
|
198 |
subst [] t |
|
53358 | 199 |
|> tap (fn u => has_call u andalso (* FIXME detect this case earlier *) |
200 |
primrec_error_eqn "recursive call not directly applied to constructor argument" t) |
|
53350 | 201 |
end; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
202 |
|
53358 | 203 |
fun build_rec_arg lthy funs_data has_call ctr_spec maybe_eqn_data = |
204 |
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
|
205 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
206 |
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
|
207 |
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
|
208 |
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
|
209 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
210 |
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
|
211 |
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
|
212 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
213 |
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
|
214 |
|> 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
|
215 |
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
|
216 |
|> 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
|
217 |
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
|
218 |
|> 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
|
219 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
220 |
fun make_direct_type T = dummyT; (* FIXME? *) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
221 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
222 |
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
|
223 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
224 |
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
|
225 |
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
|
226 |
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
|
227 |
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
|
228 |
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
|
229 |
| 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
|
230 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
231 |
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
|
232 |
|> 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
|
233 |
|> map Free |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
234 |
|> 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
|
235 |
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
|
236 |
no_calls' |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
237 |
|> 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
|
238 |
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
|
239 |
direct_calls' |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
240 |
|> 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
|
241 |
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
|
242 |
indirect_calls'; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
243 |
|
53401 | 244 |
val fun_name_ctr_pos_list = |
245 |
map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data; |
|
246 |
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
|
247 |
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
|
248 |
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
|
249 |
|
53401 | 250 |
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
|
251 |
in |
53350 | 252 |
t |
53401 | 253 |
|> subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls |
254 |
|> fold_rev lambda abstractions |
|
53350 | 255 |
end; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
256 |
|
53358 | 257 |
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
|
258 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
259 |
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
|
260 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
261 |
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
|
262 |
(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
|
263 |
|> 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
|
264 |
##> (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
|
265 |
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
|
266 |
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
|
267 |
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
|
268 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
269 |
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
|
270 |
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
|
271 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
272 |
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
|
273 |
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
|
274 |
|> sort ((op <) o pairself (#offset o fst) |> make_ord) |
53358 | 275 |
|> 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
|
276 |
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
|
277 |
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
|
278 |
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
|
279 |
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
|
280 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
281 |
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
|
282 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
283 |
(recs, ctr_poss) |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
284 |
|-> 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
|
285 |
|> Syntax.check_terms lthy |
53352 | 286 |
|> 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
|
287 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
288 |
|
53358 | 289 |
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
|
290 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
291 |
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
|
292 |
| 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
|
293 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
294 |
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
|
295 |
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
|
296 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
297 |
if n < 0 then |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
298 |
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
|
299 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
300 |
let val (f, args) = chop n args' |>> curry list_comb f' in |
53358 | 301 |
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
|
302 |
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
|
303 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
304 |
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
|
305 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
306 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
307 |
| find _ _ = []; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
308 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
309 |
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
|
310 |
|> (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
|
311 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
312 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
313 |
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
|
314 |
let |
53352 | 315 |
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
|
316 |
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
|
317 |
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
|
318 |
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
|
319 |
|> 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
|
320 |
|> 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
|
321 |
|> 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
|
322 |
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
|
323 |
|
53358 | 324 |
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
|
325 |
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
|
326 |
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
|
327 |
val callssss = funs_data |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
328 |
|> map (partition_eq ((op =) o pairself #ctr)) |
53358 | 329 |
|> 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
|
330 |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
331 |
val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') = |
53794 | 332 |
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
|
333 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
334 |
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
|
335 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
336 |
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
|
337 |
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
|
338 |
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
|
339 |
" 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
|
340 |
|
53358 | 341 |
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
|
342 |
|
53329 | 343 |
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
|
344 |
lthy = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
345 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
346 |
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
|
347 |
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
|
348 |
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
|
349 |
|> fst |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
350 |
|> 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
|
351 |
(#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
|
352 |
|> map (fn (user_eqn, num_extra_args, rec_thm) => |
53329 | 353 |
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
|
354 |
|> 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
|
355 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
356 |
val notes = |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
357 |
[(inductN, if n2m then [induct_thm] else [], []), |
53811
2967fa35d89e
set code and nitpick_simp attributes on primcorec theorems
blanchet
parents:
53797
diff
changeset
|
358 |
(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
|
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 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
|
362 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
363 |
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
|
364 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
365 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
366 |
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
|
367 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
368 |
val common_notes = |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
369 |
[(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
|
370 |
|> 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
|
371 |
|> 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
|
372 |
((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
|
373 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
374 |
lthy' |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
375 |
|> fold_map Local_Theory.define defs |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
376 |
|-> 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
|
377 |
(take actual_nn induct_thms) funs_data) |
53797 | 378 |
|> 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
|
379 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
380 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
381 |
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
|
382 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
383 |
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
|
384 |
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
|
385 |
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
|
386 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
387 |
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
|
388 |
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
|
389 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
390 |
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
|
391 |
if null eqns |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
392 |
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
|
393 |
else error ("primrec_new error:\n " ^ str ^ "\nin\n " ^ |
53822 | 394 |
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
|
395 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
396 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
397 |
|
53310 | 398 |
(* Primcorec *) |
53303
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_disc = { |
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, |
53720 | 404 |
ctr: term, |
53341 | 405 |
ctr_no: int, (*###*) |
53720 | 406 |
disc: term, |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
407 |
prems: term list, |
53822 | 408 |
auto_gen: bool, |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
409 |
user_eqn: term |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
410 |
}; |
53341 | 411 |
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
|
412 |
fun_name: string, |
53720 | 413 |
fun_T: typ, |
53401 | 414 |
fun_args: term list, |
53341 | 415 |
ctr: term, |
416 |
sel: term, |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
417 |
rhs_term: term, |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
418 |
user_eqn: term |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
419 |
}; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
420 |
datatype co_eqn_data = |
53341 | 421 |
Disc of co_eqn_data_disc | |
422 |
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
|
423 |
|
53720 | 424 |
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
|
425 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
426 |
fun find_subterm p = let (* FIXME \<exists>? *) |
53401 | 427 |
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
|
428 |
| 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
|
429 |
in f end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
430 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
431 |
val applied_fun = concl |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
432 |
|> 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
|
433 |
|> the |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
434 |
handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl; |
53720 | 435 |
val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free; |
436 |
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
|
437 |
|
53720 | 438 |
val discs = map #disc ctr_specs; |
439 |
val ctrs = map #ctr ctr_specs; |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
440 |
val not_disc = head_of concl = @{term Not}; |
53401 | 441 |
val _ = not_disc andalso length ctrs <> 2 andalso |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
442 |
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
|
443 |
val disc = find_subterm (member (op =) discs o head_of) concl; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
444 |
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
|
445 |
|> (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
|
446 |
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
|
447 |
val _ = is_some disc orelse is_some eq_ctr0 orelse |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
448 |
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
|
449 |
val ctr_no' = |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
450 |
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
|
451 |
val ctr_no = if not_disc then 1 - ctr_no' else ctr_no'; |
53720 | 452 |
val ctr = #ctr (nth ctr_specs ctr_no); |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
453 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
454 |
val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_; |
53720 | 455 |
val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default []; |
456 |
val prems = map (abstract (List.rev fun_args)) prems'; |
|
457 |
val real_prems = |
|
458 |
(if catch_all orelse sequential then maps invert_prems_disj matchedss else []) @ |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
459 |
(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
|
460 |
|
53720 | 461 |
val matchedsss' = AList.delete (op =) fun_name matchedsss |
462 |
|> 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
|
463 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
464 |
val user_eqn = |
53720 | 465 |
(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
|
466 |
|>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
467 |
|> Logic.list_implies; |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
468 |
in |
53341 | 469 |
(Disc { |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
470 |
fun_name = fun_name, |
53720 | 471 |
fun_T = fun_T, |
53401 | 472 |
fun_args = fun_args, |
53720 | 473 |
ctr = ctr, |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
474 |
ctr_no = ctr_no, |
53720 | 475 |
disc = #disc (nth ctr_specs ctr_no), |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
476 |
prems = real_prems, |
53822 | 477 |
auto_gen = catch_all, |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
478 |
user_eqn = user_eqn |
53720 | 479 |
}, matchedsss') |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
480 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
481 |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
482 |
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
|
483 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
484 |
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
|
485 |
handle TERM _ => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
486 |
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
|
487 |
val sel = head_of lhs; |
53720 | 488 |
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
|
489 |
handle TERM _ => |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
490 |
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
|
491 |
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
|
492 |
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
|
493 |
val ctr_spec = |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
494 |
if is_some of_spec |
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
495 |
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
|
496 |
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
|
497 |
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
|
498 |
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
|
499 |
in |
53341 | 500 |
Sel { |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
501 |
fun_name = fun_name, |
53720 | 502 |
fun_T = fun_T, |
53401 | 503 |
fun_args = fun_args, |
53341 | 504 |
ctr = #ctr ctr_spec, |
505 |
sel = sel, |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
506 |
rhs_term = rhs, |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
507 |
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
|
508 |
} |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
509 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
510 |
|
53720 | 511 |
fun 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
|
512 |
let |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
513 |
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
|
514 |
val fun_name = head_of lhs |> fst o dest_Free; |
53720 | 515 |
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
|
516 |
val (ctr, ctr_args) = strip_comb rhs; |
53720 | 517 |
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
|
518 |
handle Option.Option => primrec_error_eqn "not a constructor" ctr; |
53341 | 519 |
|
53720 | 520 |
val disc_imp_rhs = betapply (disc, lhs); |
521 |
val (maybe_eqn_data_disc, matchedsss') = if length ctr_specs = 1 |
|
522 |
then (NONE, matchedsss) |
|
53341 | 523 |
else apfst SOME (co_dissect_eqn_disc |
53720 | 524 |
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
|
525 |
|
53720 | 526 |
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
|
527 |
|> 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
|
528 |
|
53360 | 529 |
val _ = tracing ("reduced\n " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n \<cdot> " ^ |
53341 | 530 |
(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
|
531 |
space_implode "\n \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss)); |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
532 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
533 |
val eqns_data_sel = |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
534 |
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
|
535 |
in |
53720 | 536 |
(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
|
537 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
538 |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
539 |
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
|
540 |
let |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
541 |
val eqn = drop_All eqn' |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
542 |
handle TERM _ => primrec_error_eqn "malformed function equation" eqn'; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
543 |
val (imp_prems, imp_rhs) = Logic.strip_horn eqn |
53341 | 544 |
|> 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
|
545 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
546 |
val head = imp_rhs |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
547 |
|> 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
|
548 |
|> head_of; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
549 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
550 |
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
|
551 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
552 |
val discs = maps #ctr_specs corec_specs |> map #disc; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
553 |
val sels = maps #ctr_specs corec_specs |> maps #sels; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
554 |
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
|
555 |
in |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
556 |
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
|
557 |
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
|
558 |
member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then |
53720 | 559 |
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
|
560 |
|>> single |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
561 |
else if member (op =) sels head then |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
562 |
([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
|
563 |
else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then |
53720 | 564 |
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
|
565 |
else |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
566 |
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
|
567 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
568 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
569 |
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
|
570 |
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
|
571 |
mk_conjs prems |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
572 |
|> curry subst_bounds (List.rev fun_args) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
573 |
|> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
574 |
|> 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
|
575 |
|
53720 | 576 |
fun build_corec_arg_no_call sel_eqns sel = |
577 |
find_first (equal sel o #sel) sel_eqns |
|
578 |
|> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term) |
|
579 |
|> the_default undef_const |
|
53411 | 580 |
|> K; |
53360 | 581 |
|
53735 | 582 |
fun build_corec_args_direct_call lthy has_call sel_eqns sel = |
53360 | 583 |
let |
53411 | 584 |
val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns; |
53735 | 585 |
fun rewrite_q t = if has_call t then @{term False} else @{term True}; |
586 |
fun rewrite_g t = if has_call t then undef_const else t; |
|
587 |
fun rewrite_h t = if has_call t then HOLogic.mk_tuple (snd (strip_comb t)) else undef_const; |
|
588 |
fun massage _ NONE t = t |
|
589 |
| massage f (SOME {fun_args, rhs_term, ...}) t = |
|
590 |
massage_direct_corec_call lthy has_call f (range_type (fastype_of t)) rhs_term |
|
591 |
|> abs_tuple fun_args; |
|
53360 | 592 |
in |
53735 | 593 |
(massage rewrite_q maybe_sel_eqn, |
594 |
massage rewrite_g maybe_sel_eqn, |
|
595 |
massage rewrite_h maybe_sel_eqn) |
|
53360 | 596 |
end; |
597 |
||
53411 | 598 |
fun build_corec_arg_indirect_call lthy has_call sel_eqns sel = |
599 |
let |
|
600 |
val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns; |
|
53735 | 601 |
fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b) |
602 |
| rewrite bound_Ts U T (t as _ $ _) = |
|
53734
7613573f023a
avoid infinite loop for unapplied terms + tuning
blanchet
parents:
53733
diff
changeset
|
603 |
let val (u, vs) = strip_comb t in |
7613573f023a
avoid infinite loop for unapplied terms + tuning
blanchet
parents:
53733
diff
changeset
|
604 |
if is_Free u andalso has_call u then |
53735 | 605 |
Inr_const U T $ mk_tuple1 bound_Ts vs |
53734
7613573f023a
avoid infinite loop for unapplied terms + tuning
blanchet
parents:
53733
diff
changeset
|
606 |
else if try (fst o dest_Const) u = SOME @{const_name prod_case} then |
53735 | 607 |
list_comb (map_types (K dummyT) u, map (rewrite bound_Ts U T) vs) |
53734
7613573f023a
avoid infinite loop for unapplied terms + tuning
blanchet
parents:
53733
diff
changeset
|
608 |
else |
53735 | 609 |
list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs) |
610 |
end |
|
611 |
| rewrite _ U T t = if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t; |
|
612 |
fun massage NONE t = t |
|
613 |
| massage (SOME {fun_args, rhs_term, ...}) t = |
|
614 |
massage_indirect_corec_call lthy has_call (rewrite []) [] |
|
615 |
(range_type (fastype_of t)) rhs_term |
|
616 |
|> abs_tuple fun_args; |
|
53411 | 617 |
in |
53735 | 618 |
massage maybe_sel_eqn |
53411 | 619 |
end; |
53360 | 620 |
|
621 |
fun build_corec_args_sel lthy has_call all_sel_eqns ctr_spec = |
|
53341 | 622 |
let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in |
623 |
if null sel_eqns then I else |
|
624 |
let |
|
625 |
val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec; |
|
626 |
||
627 |
val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list; |
|
628 |
val direct_calls' = map_filter (try (apsnd (fn Direct_Corec n => n))) sel_call_list; |
|
629 |
val indirect_calls' = map_filter (try (apsnd (fn Indirect_Corec n => n))) sel_call_list; |
|
630 |
in |
|
53360 | 631 |
I |
53735 | 632 |
#> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls' |
53360 | 633 |
#> fold (fn (sel, (q, g, h)) => |
53735 | 634 |
let val (fq, fg, fh) = build_corec_args_direct_call lthy has_call sel_eqns sel in |
635 |
nth_map q fq o nth_map g fg o nth_map h fh end) direct_calls' |
|
53360 | 636 |
#> fold (fn (sel, n) => nth_map n |
53411 | 637 |
(build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls' |
53341 | 638 |
end |
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
639 |
end; |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
640 |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
641 |
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
|
642 |
let |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
643 |
val corec_specs' = take (length bs) corec_specs; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
644 |
val corecs = map #corec corec_specs'; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
645 |
val ctr_specss = map #ctr_specs corec_specs'; |
53360 | 646 |
val corec_args = hd corecs |
647 |
|> fst o split_last o binder_types o fastype_of |
|
648 |
|> map (Const o pair @{const_name undefined}) |
|
53720 | 649 |
|> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss |
53360 | 650 |
|> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss; |
53735 | 651 |
fun currys [] t = t |
652 |
| currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0)) |
|
653 |
|> fold_rev (Term.abs o pair Name.uu) Ts; |
|
53401 | 654 |
|
53360 | 655 |
val _ = tracing ("corecursor arguments:\n \<cdot> " ^ |
53411 | 656 |
space_implode "\n \<cdot> " (map (Syntax.string_of_term lthy) corec_args)); |
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), |
|
691 |
prems = maps (invert_prems 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) = |
738 |
if a orelse c = c' orelse sequential then SOME (K (mk_primcorec_assumption_tac lthy)) |
|
739 |
else if simple then SOME (K (auto_tac lthy)) |
|
740 |
else NONE; |
|
741 |
||
742 |
val _ = tracing ("exclusiveness properties:\n \<cdot> " ^ |
|
743 |
space_implode "\n \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss')); |
|
744 |
||
745 |
val exclss'' = exclss' |> map (map (fn (idx, t) => |
|
746 |
(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
|
747 |
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
|
748 |
val (obligation_idxss, obligationss) = exclss'' |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
749 |
|> 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
|
750 |
|> split_list o map split_list; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
751 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
752 |
fun prove thmss' def_thms' lthy = |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
753 |
let |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
754 |
val def_thms = map (snd o snd) def_thms'; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
755 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
756 |
val exclss' = map (op ~~) (obligation_idxss ~~ thmss'); |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
757 |
fun mk_exclsss excls n = |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
758 |
(excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1)) |
53822 | 759 |
|-> 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
|
760 |
val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs) |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
761 |
|-> 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
|
762 |
|
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
763 |
fun prove_disc {ctr_specs, ...} exclsss |
53720 | 764 |
{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
|
765 |
if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else |
53720 | 766 |
let |
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
767 |
val {disc_corec, ...} = nth ctr_specs ctr_no; |
53720 | 768 |
val k = 1 + ctr_no; |
769 |
val m = length prems; |
|
770 |
val t = |
|
771 |
list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0)) |
|
772 |
|> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*) |
|
773 |
|> HOLogic.mk_Trueprop |
|
774 |
|> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) |
|
775 |
|> curry Logic.list_all (map dest_Free fun_args); |
|
776 |
in |
|
777 |
mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss |
|
778 |
|> K |> Goal.prove lthy [] [] t |
|
779 |
|> pair (#disc (nth ctr_specs ctr_no)) |
|
780 |
|> single |
|
781 |
end; |
|
782 |
||
783 |
fun prove_sel {ctr_specs, nested_maps, nested_map_idents, nested_map_comps, ...} |
|
784 |
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
|
785 |
let |
53720 | 786 |
val (SOME ctr_spec) = find_first (equal ctr o #ctr) ctr_specs; |
787 |
val ctr_no = find_index (equal ctr o #ctr) ctr_specs; |
|
788 |
val prems = the_default (maps (invert_prems o #prems) disc_eqns) |
|
789 |
(find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems); |
|
790 |
val sel_corec = find_index (equal sel) (#sels ctr_spec) |
|
791 |
|> nth (#sel_corecs ctr_spec); |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
792 |
val k = 1 + ctr_no; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
793 |
val m = length prems; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
794 |
val t = |
53720 | 795 |
list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0)) |
796 |
|> curry betapply sel |
|
797 |
|> rpair (abstract (List.rev fun_args) rhs_term) |
|
798 |
|> HOLogic.mk_Trueprop o HOLogic.mk_eq |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
799 |
|> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) |
53720 | 800 |
|> curry Logic.list_all (map dest_Free fun_args); |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
801 |
in |
53692 | 802 |
mk_primcorec_ctr_or_sel_tac lthy def_thms sel_corec k m exclsss |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
803 |
nested_maps nested_map_idents nested_map_comps |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
804 |
|> K |> Goal.prove lthy [] [] t |
53720 | 805 |
|> pair sel |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
806 |
end; |
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
807 |
|
53791 | 808 |
fun prove_ctr disc_alist sel_alist disc_eqns sel_eqns {ctr, disc, sels, collapse, ...} = |
53720 | 809 |
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
|
810 |
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
|
811 |
andalso (warning ("no eqns for ctr " ^ Syntax.string_of_term lthy ctr); true) |
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
812 |
orelse (* don't try to prove theorems when some sel_eqns are missing *) |
53720 | 813 |
filter (equal ctr o #ctr) sel_eqns |
814 |
|> fst o finds ((op =) o apsnd #sel) sels |
|
815 |
|> exists (null o snd) |
|
816 |
andalso (warning ("sel_eqn(s) missing for ctr " ^ Syntax.string_of_term lthy ctr); true) |
|
817 |
then [] else |
|
818 |
let |
|
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
819 |
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
|
820 |
(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
|
821 |
|>> 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
|
822 |
||> 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
|
823 |
|> the o merge_options; |
53720 | 824 |
val m = length prems; |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
825 |
val t = filter (equal ctr o #ctr) sel_eqns |
53720 | 826 |
|> fst o finds ((op =) o apsnd #sel) sels |
827 |
|> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract) |
|
828 |
|> curry list_comb ctr |
|
829 |
|> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T), |
|
830 |
map Bound (length fun_args - 1 downto 0))) |
|
831 |
|> HOLogic.mk_Trueprop |
|
832 |
|> curry Logic.list_implies (map HOLogic.mk_Trueprop prems) |
|
833 |
|> curry Logic.list_all (map dest_Free fun_args); |
|
53791 | 834 |
val maybe_disc_thm = AList.lookup (op =) disc_alist disc; |
835 |
val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist); |
|
53720 | 836 |
in |
53722
e176d6d3345f
generate more theorems (e.g. for types with only one constructor)
panny
parents:
53720
diff
changeset
|
837 |
mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms |
53720 | 838 |
|> K |> Goal.prove lthy [] [] t |
839 |
|> single |
|
840 |
end; |
|
841 |
||
53791 | 842 |
val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss; |
843 |
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
|
844 |
|
53791 | 845 |
val disc_thmss = map (map snd) disc_alists; |
846 |
val sel_thmss = map (map snd) sel_alists; |
|
847 |
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
|
848 |
(map #ctr_specs corec_specs); |
53791 | 849 |
|
53744
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
850 |
val safess = map (map (K false)) ctr_thmss; (* FIXME: "true" for non-corecursive theorems *) |
53822 | 851 |
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
|
852 |
|
53795 | 853 |
fun mk_simp_thms disc_thms sel_thms ctr_thms = disc_thms @ sel_thms @ ctr_thms; |
854 |
||
855 |
val simp_thmss = map3 mk_simp_thms disc_thmss sel_thmss safe_ctr_thmss; |
|
856 |
||
53797 | 857 |
val common_name = mk_common_name fun_names; |
858 |
||
53744
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
859 |
val anonymous_notes = |
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
860 |
[(flat safe_ctr_thmss, simp_attrs)] |
9db52ae90009
setting the stage for safe constructor simp rules
blanchet
parents:
53743
diff
changeset
|
861 |
|> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); |
53791 | 862 |
|
863 |
val notes = |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
864 |
[(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
|
865 |
(codeN, ctr_thmss(*FIXME*), code_nitpick_simp_attrs), |
53797 | 866 |
(ctrN, ctr_thmss, []), |
53791 | 867 |
(discN, disc_thmss, simp_attrs), |
53795 | 868 |
(selN, sel_thmss, simp_attrs), |
53797 | 869 |
(simpsN, simp_thmss, []), |
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
870 |
(strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])] |
53791 | 871 |
|> maps (fn (thmN, thmss, attrs) => |
872 |
map2 (fn fun_name => fn thms => |
|
873 |
((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
|
874 |
fun_names (take actual_nn thmss)) |
53791 | 875 |
|> filter_out (null o fst o hd o snd); |
53797 | 876 |
|
877 |
val common_notes = |
|
53830
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
878 |
[(coinductN, if n2m then [coinduct_thm] else [], []), |
ed2eb7df2aac
don't note more induction principles than there are functions + tuning
blanchet
parents:
53822
diff
changeset
|
879 |
(strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])] |
53797 | 880 |
|> filter_out (null o #2) |
881 |
|> map (fn (thmN, thms, attrs) => |
|
882 |
((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])])); |
|
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
883 |
in |
53797 | 884 |
lthy |> Local_Theory.notes (anonymous_notes @ notes @ common_notes) |> snd |
53654
8b9ea4420f81
prove simp theorems for newly generated definitions
panny
parents:
53411
diff
changeset
|
885 |
end; |
53822 | 886 |
|
887 |
fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss'; |
|
888 |
||
889 |
val _ = if not simple orelse forall null obligationss then () else |
|
890 |
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
|
891 |
in |
53822 | 892 |
if simple then |
893 |
lthy' |
|
894 |
|> after_qed (map (fn [] => []) obligationss) |
|
895 |
|> pair NONE o SOME |
|
896 |
else |
|
897 |
lthy' |
|
898 |
|> Proof.theorem NONE after_qed obligationss |
|
899 |
|> Proof.refine (Method.primitive_text I) |
|
900 |
|> Seq.hd |
|
901 |
|> rpair NONE o SOME |
|
902 |
end; |
|
53303
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
903 |
|
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
904 |
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
|
905 |
let |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
906 |
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
|
907 |
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
|
908 |
in |
53831
80423b9080cf
support "of" syntax to disambiguate selector equations
panny
parents:
53830
diff
changeset
|
909 |
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
|
910 |
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
|
911 |
end |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
912 |
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
|
913 |
if null eqns |
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
914 |
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
|
915 |
else error ("primcorec error:\n " ^ str ^ "\nin\n " ^ |
53822 | 916 |
space_implode "\n " (map (quote o Syntax.string_of_term lthy) eqns)); |
917 |
||
918 |
val add_primcorecursive_cmd = (the o fst) ooo add_primcorec_ursive_cmd false; |
|
919 |
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
|
920 |
|
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff
changeset
|
921 |
end; |