author | wenzelm |
Wed, 12 Mar 2025 11:39:00 +0100 | |
changeset 82265 | 4b875a4c83b0 |
parent 81543 | fa37ee54644c |
permissions | -rw-r--r-- |
54701 | 1 |
(* Title: HOL/Tools/Ctr_Sugar/case_translation.ML |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
2 |
Author: Konrad Slind, Cambridge University Computer Laboratory |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
3 |
Author: Stefan Berghofer, TU Muenchen |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
4 |
Author: Dmitriy Traytel, TU Muenchen |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
5 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
6 |
Nested case expressions via a generic data slot for case combinators and constructors. |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
7 |
*) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
8 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
9 |
signature CASE_TRANSLATION = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
10 |
sig |
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
11 |
val indexify_names: string list -> string list |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
12 |
val make_tnames: typ list -> string list |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
13 |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
14 |
datatype config = Error | Warning | Quiet |
51678
1e33b81c328a
allow redundant cases in the list comprehension translation
traytel
parents:
51677
diff
changeset
|
15 |
val case_tr: bool -> Proof.context -> term list -> term |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
16 |
val lookup_by_constr: Proof.context -> string * typ -> (term * term list) option |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
17 |
val lookup_by_constr_permissive: Proof.context -> string * typ -> (term * term list) option |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
18 |
val lookup_by_case: Proof.context -> string -> (term * term list) option |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
19 |
val make_case: Proof.context -> config -> Name.context -> term -> (term * term) list -> term |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
20 |
val print_case_translations: Proof.context -> unit |
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
21 |
val strip_case: Proof.context -> bool -> term -> (term * (term * term) list) option |
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
22 |
val strip_case_full: Proof.context -> bool -> term -> term |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
23 |
val show_cases: bool Config.T |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
24 |
val register: term -> term list -> Context.generic -> Context.generic |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
25 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
26 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
27 |
structure Case_Translation: CASE_TRANSLATION = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
28 |
struct |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
29 |
|
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
30 |
(** general utilities **) |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
31 |
|
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
32 |
fun indexify_names names = |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
33 |
let |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
34 |
fun index (x :: xs) tab = |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
35 |
(case AList.lookup (op =) tab x of |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
36 |
NONE => |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
37 |
if member (op =) xs x |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
38 |
then (x ^ "1") :: index xs ((x, 2) :: tab) |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
39 |
else x :: index xs tab |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
40 |
| SOME i => (x ^ string_of_int i) :: index xs ((x, i + 1) :: tab)) |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
41 |
| index [] _ = []; |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
42 |
in index names [] end; |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
43 |
|
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
44 |
fun make_tnames Ts = |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
45 |
let |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
46 |
fun type_name (TFree (name, _)) = unprefix "'" name |
71224
54a7ad860a76
made internal name generation in case expressions more robust
traytel
parents:
69593
diff
changeset
|
47 |
| type_name (TVar ((name, idx), _)) = |
54a7ad860a76
made internal name generation in case expressions more robust
traytel
parents:
69593
diff
changeset
|
48 |
unprefix "'" name ^ (if idx = 0 then "" else string_of_int idx) |
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
49 |
| type_name (Type (name, _)) = |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
50 |
let val name' = Long_Name.base_name name |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
51 |
in if Symbol_Pos.is_identifier name' then name' else "x" end; |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
52 |
in indexify_names (map type_name Ts) end; |
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
53 |
|
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
54 |
|
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52705
diff
changeset
|
55 |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
56 |
(** data management **) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
57 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
58 |
datatype data = Data of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
59 |
{constrs: (string * (term * term list)) list Symtab.table, |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
60 |
cases: (term * term list) Symtab.table}; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
61 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
62 |
fun make_data (constrs, cases) = Data {constrs = constrs, cases = cases}; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
63 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
64 |
structure Data = Generic_Data |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
65 |
( |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
66 |
type T = data; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
67 |
val empty = make_data (Symtab.empty, Symtab.empty); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
68 |
fun merge |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
69 |
(Data {constrs = constrs1, cases = cases1}, |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
70 |
Data {constrs = constrs2, cases = cases2}) = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
71 |
make_data |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
72 |
(Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2), |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
73 |
Symtab.merge (K true) (cases1, cases2)); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
74 |
); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
75 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
76 |
fun map_data f = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
77 |
Data.map (fn Data {constrs, cases} => make_data (f (constrs, cases))); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
78 |
fun map_constrs f = map_data (fn (constrs, cases) => (f constrs, cases)); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
79 |
fun map_cases f = map_data (fn (constrs, cases) => (constrs, f cases)); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
80 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
81 |
val rep_data = (fn Data args => args) o Data.get o Context.Proof; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
82 |
|
51684 | 83 |
fun T_of_data (comb, constrs : term list) = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
84 |
fastype_of comb |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
85 |
|> funpow (length constrs) range_type |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
86 |
|> domain_type; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
87 |
|
80634
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
wenzelm
parents:
74561
diff
changeset
|
88 |
val Tname_of_data = dest_Type_name o T_of_data; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
89 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
90 |
val constrs_of = #constrs o rep_data; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
91 |
val cases_of = #cases o rep_data; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
92 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
93 |
fun lookup_by_constr ctxt (c, T) = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
94 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
95 |
val tab = Symtab.lookup_list (constrs_of ctxt) c; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
96 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
97 |
(case body_type T of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
98 |
Type (tyco, _) => AList.lookup (op =) tab tyco |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
99 |
| _ => NONE) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
100 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
101 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
102 |
fun lookup_by_constr_permissive ctxt (c, T) = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
103 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
104 |
val tab = Symtab.lookup_list (constrs_of ctxt) c; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
105 |
val hint = (case body_type T of Type (tyco, _) => SOME tyco | _ => NONE); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
106 |
val default = if null tab then NONE else SOME (snd (List.last tab)); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
107 |
(*conservative wrt. overloaded constructors*) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
108 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
109 |
(case hint of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
110 |
NONE => default |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
111 |
| SOME tyco => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
112 |
(case AList.lookup (op =) tab tyco of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
113 |
NONE => default (*permissive*) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
114 |
| SOME info => SOME info)) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
115 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
116 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
117 |
val lookup_by_case = Symtab.lookup o cases_of; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
118 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
119 |
|
52154 | 120 |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
121 |
(** installation **) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
122 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
123 |
fun case_error s = error ("Error in case expression:\n" ^ s); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
124 |
|
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
125 |
val name_of = try dest_Const_name; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
126 |
|
52154 | 127 |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
128 |
(* parse translation *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
129 |
|
69593 | 130 |
fun constrain_Abs tT t = Syntax.const \<^syntax_const>\<open>_constrainAbs\<close> $ t $ tT; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
131 |
|
51678
1e33b81c328a
allow redundant cases in the list comprehension translation
traytel
parents:
51677
diff
changeset
|
132 |
fun case_tr err ctxt [t, u] = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
133 |
let |
52157 | 134 |
val consts = Proof_Context.consts_of ctxt; |
135 |
fun is_const s = can (Consts.the_constraint consts) (Consts.intern consts s); |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
136 |
|
52158
d5fa81343322
more robust variant_free: avoid clash with consts name space (e.g. consts "x", "xa", etc.);
wenzelm
parents:
52157
diff
changeset
|
137 |
fun variant_free x used = |
d5fa81343322
more robust variant_free: avoid clash with consts name space (e.g. consts "x", "xa", etc.);
wenzelm
parents:
52157
diff
changeset
|
138 |
let val (x', used') = Name.variant x used |
d5fa81343322
more robust variant_free: avoid clash with consts name space (e.g. consts "x", "xa", etc.);
wenzelm
parents:
52157
diff
changeset
|
139 |
in if is_const x' then variant_free x' used' else (x', used') end; |
d5fa81343322
more robust variant_free: avoid clash with consts name space (e.g. consts "x", "xa", etc.);
wenzelm
parents:
52157
diff
changeset
|
140 |
|
52157 | 141 |
fun abs p tTs t = |
69593 | 142 |
Syntax.const \<^const_syntax>\<open>case_abs\<close> $ |
52157 | 143 |
fold constrain_Abs tTs (absfree p t); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
144 |
|
69593 | 145 |
fun abs_pat (Const (\<^syntax_const>\<open>_constrain\<close>, _) $ t $ tT) tTs = |
51693 | 146 |
abs_pat t (tT :: tTs) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
147 |
| abs_pat (Free (p as (x, _))) tTs = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
148 |
if is_const x then I else abs p tTs |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
149 |
| abs_pat (t $ u) _ = abs_pat u [] #> abs_pat t [] |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
150 |
| abs_pat _ _ = I; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
151 |
|
51680
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
152 |
(* replace occurrences of dummy_pattern by distinct variables *) |
69593 | 153 |
fun replace_dummies (Const (\<^const_syntax>\<open>Pure.dummy_pattern\<close>, T)) used = |
52158
d5fa81343322
more robust variant_free: avoid clash with consts name space (e.g. consts "x", "xa", etc.);
wenzelm
parents:
52157
diff
changeset
|
154 |
let val (x, used') = variant_free "x" used |
51680
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
155 |
in (Free (x, T), used') end |
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
156 |
| replace_dummies (t $ u) used = |
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
157 |
let |
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
158 |
val (t', used') = replace_dummies t used; |
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
159 |
val (u', used'') = replace_dummies u used'; |
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
160 |
in (t' $ u', used'') end |
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
161 |
| replace_dummies t used = (t, used); |
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
162 |
|
69593 | 163 |
fun dest_case1 (t as Const (\<^syntax_const>\<open>_case1\<close>, _) $ l $ r) = |
81506
f76a5849b570
tuned: more standard Name.build_context, although that is a bit longer;
wenzelm
parents:
81505
diff
changeset
|
164 |
let val (l', _) = replace_dummies l (Name.build_context (Term.declare_free_names t)) in |
52154 | 165 |
abs_pat l' [] |
69593 | 166 |
(Syntax.const \<^const_syntax>\<open>case_elem\<close> $ Term_Position.strip_positions l' $ r) |
51680
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
167 |
end |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
168 |
| dest_case1 _ = case_error "dest_case1"; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
169 |
|
69593 | 170 |
fun dest_case2 (Const (\<^syntax_const>\<open>_case2\<close>, _) $ t $ u) = t :: dest_case2 u |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
171 |
| dest_case2 t = [t]; |
51678
1e33b81c328a
allow redundant cases in the list comprehension translation
traytel
parents:
51677
diff
changeset
|
172 |
|
69593 | 173 |
val errt = Syntax.const (if err then \<^const_syntax>\<open>True\<close> else \<^const_syntax>\<open>False\<close>); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
174 |
in |
69593 | 175 |
Syntax.const \<^const_syntax>\<open>case_guard\<close> $ errt $ t $ |
52154 | 176 |
(fold_rev |
69593 | 177 |
(fn t => fn u => Syntax.const \<^const_syntax>\<open>case_cons\<close> $ dest_case1 t $ u) |
52154 | 178 |
(dest_case2 u) |
69593 | 179 |
(Syntax.const \<^const_syntax>\<open>case_nil\<close>)) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
180 |
end |
51678
1e33b81c328a
allow redundant cases in the list comprehension translation
traytel
parents:
51677
diff
changeset
|
181 |
| case_tr _ _ _ = case_error "case_tr"; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
182 |
|
69593 | 183 |
val _ = Theory.setup (Sign.parse_translation [(\<^syntax_const>\<open>_case_syntax\<close>, case_tr true)]); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
184 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
185 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
186 |
(* print translation *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
187 |
|
51751
cf039b3c42a7
slightly more aggressive syntax translation for printing case expressions
traytel
parents:
51693
diff
changeset
|
188 |
fun case_tr' (_ :: x :: t :: ts) = |
52159 | 189 |
let |
69593 | 190 |
fun mk_clause (Const (\<^const_syntax>\<open>case_abs\<close>, _) $ Abs (s, T, t)) xs used = |
52159 | 191 |
let val (s', used') = Name.variant s used |
192 |
in mk_clause t ((s', T) :: xs) used' end |
|
69593 | 193 |
| mk_clause (Const (\<^const_syntax>\<open>case_elem\<close>, _) $ pat $ rhs) xs _ = |
194 |
Syntax.const \<^syntax_const>\<open>_case1\<close> $ |
|
52159 | 195 |
subst_bounds (map Syntax_Trans.mark_bound_abs xs, pat) $ |
196 |
subst_bounds (map Syntax_Trans.mark_bound_body xs, rhs) |
|
197 |
| mk_clause _ _ _ = raise Match; |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
198 |
|
69593 | 199 |
fun mk_clauses (Const (\<^const_syntax>\<open>case_nil\<close>, _)) = [] |
200 |
| mk_clauses (Const (\<^const_syntax>\<open>case_cons\<close>, _) $ t $ u) = |
|
81506
f76a5849b570
tuned: more standard Name.build_context, although that is a bit longer;
wenzelm
parents:
81505
diff
changeset
|
201 |
mk_clause t [] (Name.build_context (Term.declare_free_names t)) :: mk_clauses u |
52159 | 202 |
| mk_clauses _ = raise Match; |
203 |
in |
|
69593 | 204 |
list_comb (Syntax.const \<^syntax_const>\<open>_case_syntax\<close> $ x $ |
205 |
foldr1 (fn (t, u) => Syntax.const \<^syntax_const>\<open>_case2\<close> $ t $ u) |
|
52159 | 206 |
(mk_clauses t), ts) |
207 |
end |
|
208 |
| case_tr' _ = raise Match; |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
209 |
|
69593 | 210 |
val _ = Theory.setup (Sign.print_translation [(\<^const_syntax>\<open>case_guard\<close>, K case_tr')]); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
211 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
212 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
213 |
(* declarations *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
214 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
215 |
fun register raw_case_comb raw_constrs context = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
216 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
217 |
val ctxt = Context.proof_of context; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
218 |
val case_comb = singleton (Variable.polymorphic ctxt) raw_case_comb; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
219 |
val constrs = Variable.polymorphic ctxt raw_constrs; |
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
220 |
val case_key = dest_Const_name case_comb; |
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
221 |
val constr_keys = map dest_Const_name constrs; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
222 |
val data = (case_comb, constrs); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
223 |
val Tname = Tname_of_data data; |
55444
ec73f81e49e7
iteration n in the 'default' vs. 'update_new' vs. 'update' saga -- 'update' makes sense now that we honor the canonical order on 'merge' (as opposed to raising 'DUP')
blanchet
parents:
55392
diff
changeset
|
224 |
val update_constrs = fold (fn key => Symtab.cons_list (key, (Tname, data))) constr_keys; |
ec73f81e49e7
iteration n in the 'default' vs. 'update_new' vs. 'update' saga -- 'update' makes sense now that we honor the canonical order on 'merge' (as opposed to raising 'DUP')
blanchet
parents:
55392
diff
changeset
|
225 |
val update_cases = Symtab.update (case_key, data); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
226 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
227 |
context |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
228 |
|> map_constrs update_constrs |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
229 |
|> map_cases update_cases |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
230 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
231 |
|
55971 | 232 |
val _ = Theory.setup |
69593 | 233 |
(Attrib.setup \<^binding>\<open>case_translation\<close> |
55971 | 234 |
(Args.term -- Scan.repeat1 Args.term >> |
235 |
(fn (t, ts) => Thm.declaration_attribute (K (register t ts)))) |
|
236 |
"declaration of case combinators and constructors"); |
|
237 |
||
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
238 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
239 |
(* (Un)check phases *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
240 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
241 |
datatype config = Error | Warning | Quiet; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
242 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
243 |
exception CASE_ERROR of string * int; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
244 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
245 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
246 |
(*Each pattern carries with it a tag i, which denotes the clause it |
52154 | 247 |
came from. i = ~1 indicates that the clause was added by pattern |
248 |
completion.*) |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
249 |
|
52159 | 250 |
fun add_row_used ((prfx, pats), (tm, _)) = |
81505 | 251 |
fold Term.declare_free_names (tm :: pats @ map Free prfx); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
252 |
|
52154 | 253 |
(*try to preserve names given by user*) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
254 |
fun default_name "" (Free (name', _)) = name' |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
255 |
| default_name name _ = name; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
256 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
257 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
258 |
(*Produce an instance of a constructor, plus fresh variables for its arguments.*) |
51675 | 259 |
fun fresh_constr colty used c = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
260 |
let |
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
80634
diff
changeset
|
261 |
val T = dest_Const_type c; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
262 |
val Ts = binder_types T; |
81521 | 263 |
val names = Name.variants used (make_tnames Ts); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
264 |
val ty = body_type T; |
51675 | 265 |
val ty_theta = Type.raw_match (ty, colty) Vartab.empty |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
266 |
handle Type.TYPE_MATCH => raise CASE_ERROR ("type mismatch", ~1); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
267 |
val c' = Envir.subst_term_types ty_theta c; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
268 |
val gvars = map (Envir.subst_term_types ty_theta o Free) (names ~~ Ts); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
269 |
in (c', gvars) end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
270 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
271 |
(*Go through a list of rows and pick out the ones beginning with a |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
272 |
pattern with constructor = name.*) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
273 |
fun mk_group (name, T) rows = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
274 |
let val k = length (binder_types T) in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
275 |
fold (fn (row as ((prfx, p :: ps), rhs as (_, i))) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
276 |
fn ((in_group, not_in_group), names) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
277 |
(case strip_comb p of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
278 |
(Const (name', _), args) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
279 |
if name = name' then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
280 |
if length args = k then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
281 |
((((prfx, args @ ps), rhs) :: in_group, not_in_group), |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
282 |
map2 default_name names args) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
283 |
else raise CASE_ERROR ("Wrong number of arguments for constructor " ^ quote name, i) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
284 |
else ((in_group, row :: not_in_group), names) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
285 |
| _ => raise CASE_ERROR ("Not a constructor pattern", i))) |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58893
diff
changeset
|
286 |
rows (([], []), replicate k "") |>> apply2 rev |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
287 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
288 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
289 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
290 |
(* Partitioning *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
291 |
|
51675 | 292 |
fun partition _ _ _ _ [] = raise CASE_ERROR ("partition: no rows", ~1) |
52154 | 293 |
| partition used constructors colty res_ty (rows as (((prfx, _ :: ps), _) :: _)) = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
294 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
295 |
fun part [] [] = [] |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
296 |
| part [] ((_, (_, i)) :: _) = raise CASE_ERROR ("Not a constructor pattern", i) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
297 |
| part (c :: cs) rows = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
298 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
299 |
val ((in_group, not_in_group), names) = mk_group (dest_Const c) rows; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
300 |
val used' = fold add_row_used in_group used; |
51675 | 301 |
val (c', gvars) = fresh_constr colty used' c; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
302 |
val in_group' = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
303 |
if null in_group (* Constructor not given *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
304 |
then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
305 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
306 |
val Ts = map fastype_of ps; |
81521 | 307 |
val xs = |
308 |
Name.variants (fold Term.declare_free_names gvars used') |
|
309 |
(replicate (length ps) "x"); |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
310 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
311 |
[((prfx, gvars @ map Free (xs ~~ Ts)), |
69593 | 312 |
(Const (\<^const_name>\<open>undefined\<close>, res_ty), ~1))] |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
313 |
end |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
314 |
else in_group; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
315 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
316 |
{constructor = c', |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
317 |
new_formals = gvars, |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
318 |
names = names, |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
319 |
group = in_group'} :: part cs not_in_group |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
320 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
321 |
in part constructors rows end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
322 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
323 |
fun v_to_prfx (prfx, Free v :: pats) = (v :: prfx, pats) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
324 |
| v_to_prfx _ = raise CASE_ERROR ("mk_case: v_to_prfx", ~1); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
325 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
326 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
327 |
(* Translation of pattern terms into nested case expressions. *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
328 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
329 |
fun mk_case ctxt used range_ty = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
330 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
331 |
val get_info = lookup_by_constr_permissive ctxt; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
332 |
|
52159 | 333 |
fun expand _ _ _ ((_, []), _) = raise CASE_ERROR ("mk_case: expand", ~1) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
334 |
| expand constructors used ty (row as ((prfx, p :: ps), (rhs, tag))) = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
335 |
if is_Free p then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
336 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
337 |
val used' = add_row_used row used; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
338 |
fun expnd c = |
51675 | 339 |
let val capp = list_comb (fresh_constr ty used' c) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
340 |
in ((prfx, capp :: ps), (subst_free [(p, capp)] rhs, tag)) end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
341 |
in map expnd constructors end |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
342 |
else [row]; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
343 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
344 |
val (name, _) = Name.variant "a" used; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
345 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
346 |
fun mk _ [] = raise CASE_ERROR ("no rows", ~1) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
347 |
| mk [] (((_, []), (tm, tag)) :: _) = ([tag], tm) (* Done *) |
52159 | 348 |
| mk path ((row as ((_, [Free _]), _)) :: _ :: _) = mk path [row] |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
349 |
| mk (u :: us) (rows as ((_, _ :: _), _) :: _) = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
350 |
let val col0 = map (fn ((_, p :: _), (_, i)) => (p, i)) rows in |
52154 | 351 |
(case Option.map (apfst head_of) (find_first (not o is_Free o fst) col0) of |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
352 |
NONE => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
353 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
354 |
val rows' = map (fn ((v, _), row) => row ||> |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
355 |
apfst (subst_free [(v, u)]) |>> v_to_prfx) (col0 ~~ rows); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
356 |
in mk us rows' end |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
357 |
| SOME (Const (cname, cT), i) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
358 |
(case get_info (cname, cT) of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
359 |
NONE => raise CASE_ERROR ("Not a datatype constructor: " ^ quote cname, i) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
360 |
| SOME (case_comb, constructors) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
361 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
362 |
val pty = body_type cT; |
81505 | 363 |
val used' = fold Term.declare_free_names us used; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
364 |
val nrows = maps (expand constructors used' pty) rows; |
52154 | 365 |
val subproblems = partition used' constructors pty range_ty nrows; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
366 |
val (pat_rect, dtrees) = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
367 |
split_list (map (fn {new_formals, group, ...} => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
368 |
mk (new_formals @ us) group) subproblems); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
369 |
val case_functions = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
370 |
map2 (fn {new_formals, names, ...} => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
371 |
fold_rev (fn (x as Free (_, T), s) => fn t => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
372 |
Abs (if s = "" then name else s, T, abstract_over (x, t))) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
373 |
(new_formals ~~ names)) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
374 |
subproblems dtrees; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
375 |
val types = map fastype_of (case_functions @ [u]); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
376 |
val case_const = Const (name_of case_comb |> the, types ---> range_ty); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
377 |
val tree = list_comb (case_const, case_functions @ [u]); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
378 |
in (flat pat_rect, tree) end) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
379 |
| SOME (t, i) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
380 |
raise CASE_ERROR ("Not a datatype constructor: " ^ Syntax.string_of_term ctxt t, i)) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
381 |
end |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
382 |
| mk _ _ = raise CASE_ERROR ("Malformed row matrix", ~1) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
383 |
in mk end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
384 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
385 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
386 |
(*Repeated variable occurrences in a pattern are not allowed.*) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
387 |
fun no_repeat_vars ctxt pat = fold_aterms |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
388 |
(fn x as Free (s, _) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
389 |
(fn xs => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
390 |
if member op aconv xs x then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
391 |
case_error (quote s ^ " occurs repeatedly in the pattern " ^ |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
392 |
quote (Syntax.string_of_term ctxt pat)) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
393 |
else x :: xs) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
394 |
| _ => I) pat []; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
395 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
396 |
fun make_case ctxt config used x clauses = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
397 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
398 |
fun string_of_clause (pat, rhs) = |
52705
c12602c1b13b
do not apply string_of_term to dummy-typed syntactic constants (ensures that uncheck phases work on well-typed terms)
traytel
parents:
52690
diff
changeset
|
399 |
Syntax.unparse_term ctxt |
69593 | 400 |
(Term.list_comb (Syntax.const \<^syntax_const>\<open>_case1\<close>, |
52705
c12602c1b13b
do not apply string_of_term to dummy-typed syntactic constants (ensures that uncheck phases work on well-typed terms)
traytel
parents:
52690
diff
changeset
|
401 |
Syntax.uncheck_terms ctxt [pat, rhs])) |
c12602c1b13b
do not apply string_of_term to dummy-typed syntactic constants (ensures that uncheck phases work on well-typed terms)
traytel
parents:
52690
diff
changeset
|
402 |
|> Pretty.string_of; |
52690
2fa3110539a5
more opportunistic string_of_clause, to make double-sure its Syntax.string_of_term uncheck phase does not crash, and thus bomb ambiguous input via failed composition of warning (e.g. HOL/Imperative_HOL/ex/Linked_List.thy: lemma merge_simps);
wenzelm
parents:
52685
diff
changeset
|
403 |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
404 |
val _ = map (no_repeat_vars ctxt o fst) clauses; |
51680
8b8cd5a527bc
Handle dummy patterns in parse translation rather than check phase
berghofe
parents:
51679
diff
changeset
|
405 |
val rows = map_index (fn (i, (pat, rhs)) => (([], [pat]), (rhs, i))) clauses; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
406 |
val rangeT = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
407 |
(case distinct (op =) (map (fastype_of o snd) clauses) of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
408 |
[] => case_error "no clauses given" |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
409 |
| [T] => T |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
410 |
| _ => case_error "all cases must have the same result type"); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
411 |
val used' = fold add_row_used rows used; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
412 |
val (tags, case_tm) = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
413 |
mk_case ctxt used' rangeT [x] rows |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
414 |
handle CASE_ERROR (msg, i) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
415 |
case_error |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
416 |
(msg ^ (if i < 0 then "" else "\nIn clause\n" ^ string_of_clause (nth clauses i))); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
417 |
val _ = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
418 |
(case subtract (op =) tags (map (snd o snd) rows) of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
419 |
[] => () |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
420 |
| is => |
52685
554d684d8520
non-strict "Quiet" mode, to avoid odd crash of code_pred/after_qed due to ill-typed use of string_of_clause (e.g. in JinjaThreads/Common/Decl.thy of AFP/d1bb08f92ce5);
wenzelm
parents:
52285
diff
changeset
|
421 |
if config = Quiet then () |
554d684d8520
non-strict "Quiet" mode, to avoid odd crash of code_pred/after_qed due to ill-typed use of string_of_clause (e.g. in JinjaThreads/Common/Decl.thy of AFP/d1bb08f92ce5);
wenzelm
parents:
52285
diff
changeset
|
422 |
else |
52690
2fa3110539a5
more opportunistic string_of_clause, to make double-sure its Syntax.string_of_term uncheck phase does not crash, and thus bomb ambiguous input via failed composition of warning (e.g. HOL/Imperative_HOL/ex/Linked_List.thy: lemma merge_simps);
wenzelm
parents:
52685
diff
changeset
|
423 |
(if config = Error then case_error else warning (*FIXME lack of syntactic context!?*)) |
52685
554d684d8520
non-strict "Quiet" mode, to avoid odd crash of code_pred/after_qed due to ill-typed use of string_of_clause (e.g. in JinjaThreads/Common/Decl.thy of AFP/d1bb08f92ce5);
wenzelm
parents:
52285
diff
changeset
|
424 |
("The following clauses are redundant (covered by preceding clauses):\n" ^ |
554d684d8520
non-strict "Quiet" mode, to avoid odd crash of code_pred/after_qed due to ill-typed use of string_of_clause (e.g. in JinjaThreads/Common/Decl.thy of AFP/d1bb08f92ce5);
wenzelm
parents:
52285
diff
changeset
|
425 |
cat_lines (map (string_of_clause o nth clauses) is))); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
426 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
427 |
case_tm |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
428 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
429 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
430 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
431 |
(* term check *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
432 |
|
69593 | 433 |
fun decode_clause (Const (\<^const_name>\<open>case_abs\<close>, _) $ Abs (s, T, t)) xs used = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
434 |
let val (s', used') = Name.variant s used |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
435 |
in decode_clause t (Free (s', T) :: xs) used' end |
69593 | 436 |
| decode_clause (Const (\<^const_name>\<open>case_elem\<close>, _) $ t $ u) xs _ = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
437 |
(subst_bounds (xs, t), subst_bounds (xs, u)) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
438 |
| decode_clause _ _ _ = case_error "decode_clause"; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
439 |
|
69593 | 440 |
fun decode_cases (Const (\<^const_name>\<open>case_nil\<close>, _)) = [] |
441 |
| decode_cases (Const (\<^const_name>\<open>case_cons\<close>, _) $ t $ u) = |
|
81506
f76a5849b570
tuned: more standard Name.build_context, although that is a bit longer;
wenzelm
parents:
81505
diff
changeset
|
442 |
decode_clause t [] (Name.build_context (Term.declare_free_names t)) :: decode_cases u |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
443 |
| decode_cases _ = case_error "decode_cases"; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
444 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
445 |
fun check_case ctxt = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
446 |
let |
69593 | 447 |
fun decode_case (Const (\<^const_name>\<open>case_guard\<close>, _) $ b $ u $ t) = |
448 |
make_case ctxt (if b = \<^term>\<open>True\<close> then Error else Warning) |
|
51678
1e33b81c328a
allow redundant cases in the list comprehension translation
traytel
parents:
51677
diff
changeset
|
449 |
Name.context (decode_case u) (decode_cases t) |
51676 | 450 |
| decode_case (t $ u) = decode_case t $ decode_case u |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
71224
diff
changeset
|
451 |
| decode_case (t as Abs _) = |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
71224
diff
changeset
|
452 |
let val (v, t') = Term.dest_abs_global t; |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
71224
diff
changeset
|
453 |
in Term.absfree v (decode_case t') end |
51676 | 454 |
| decode_case t = t; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
455 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
456 |
map decode_case |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
457 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
458 |
|
55971 | 459 |
val _ = Context.>> (Syntax_Phases.term_check 1 "case" check_case); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
460 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
461 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
462 |
(* Pretty printing of nested case expressions *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
463 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
464 |
(* destruct one level of pattern matching *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
465 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
466 |
fun dest_case ctxt d used t = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
467 |
(case apfst name_of (strip_comb t) of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
468 |
(SOME cname, ts as _ :: _) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
469 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
470 |
val (fs, x) = split_last ts; |
61004
1dd6669ff612
don't use types that come from the database---they are inconsistent with the ones occurring in the terms
traytel
parents:
60924
diff
changeset
|
471 |
fun strip_abs i t = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
472 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
473 |
val zs = strip_abs_vars t; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
474 |
val j = length zs; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
475 |
val (xs, ys) = |
61004
1dd6669ff612
don't use types that come from the database---they are inconsistent with the ones occurring in the terms
traytel
parents:
60924
diff
changeset
|
476 |
if j < i then (zs @ |
1dd6669ff612
don't use types that come from the database---they are inconsistent with the ones occurring in the terms
traytel
parents:
60924
diff
changeset
|
477 |
map (pair "x") (drop j (take i (binder_types (fastype_of t)))), []) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
478 |
else chop i zs; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
479 |
val u = fold_rev Term.abs ys (strip_abs_body t); |
81543
fa37ee54644c
clarified printing of consts: rename apart from all bounds, and thus avoid old Term.declare_free_names with its adhoc policy ("as they are printed");
wenzelm
parents:
81521
diff
changeset
|
480 |
val xs' = map Free (Name.variant_names (Term.declare_free_names u used) xs); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
481 |
val (xs1, xs2) = chop j xs' |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
482 |
in (xs', list_comb (subst_bounds (rev xs1, u), xs2)) end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
483 |
fun is_dependent i t = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
484 |
let val k = length (strip_abs_vars t) - i |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
485 |
in k < 0 orelse exists (fn j => j >= k) (loose_bnos (strip_abs_body t)) end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
486 |
fun count_cases (_, _, true) = I |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
487 |
| count_cases (c, (_, body), false) = AList.map_default op aconv (body, []) (cons c); |
69593 | 488 |
val is_undefined = name_of #> equal (SOME \<^const_name>\<open>undefined\<close>); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
489 |
fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
490 |
val get_info = lookup_by_case ctxt; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
491 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
492 |
(case get_info cname of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
493 |
SOME (_, constructors) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
494 |
if length fs = length constructors then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
495 |
let |
64961 | 496 |
val R = fastype_of x; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
497 |
val cases = map (fn (Const (s, U), t) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
498 |
let |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
499 |
val Us = binder_types U; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
500 |
val k = length Us; |
61004
1dd6669ff612
don't use types that come from the database---they are inconsistent with the ones occurring in the terms
traytel
parents:
60924
diff
changeset
|
501 |
val p as (xs, _) = strip_abs k t; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
502 |
in |
64961 | 503 |
(Const (s, map fastype_of xs ---> R), p, is_dependent k t) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
504 |
end) (constructors ~~ fs); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
505 |
val cases' = |
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
58893
diff
changeset
|
506 |
sort (int_ord o swap o apply2 (length o snd)) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
507 |
(fold_rev count_cases cases []); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
508 |
val dummy = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
509 |
if d then Term.dummy_pattern R |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
510 |
else Free (Name.variant "x" used |> fst, R); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
511 |
in |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
512 |
SOME (x, |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
513 |
map mk_case |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
514 |
(case find_first (is_undefined o fst) cases' of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
515 |
SOME (_, cs) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
516 |
if length cs = length constructors then [hd cases] |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
517 |
else filter_out (fn (_, (_, body), _) => is_undefined body) cases |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
518 |
| NONE => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
519 |
(case cases' of |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
520 |
[] => cases |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
521 |
| (default, cs) :: _ => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
522 |
if length cs = 1 then cases |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
523 |
else if length cs = length constructors then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
524 |
[hd cases, (dummy, ([], default), false)] |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
525 |
else |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
526 |
filter_out (fn (c, _, _) => member op aconv cs c) cases @ |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
527 |
[(dummy, ([], default), false)]))) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
528 |
end |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
529 |
else NONE |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
530 |
| _ => NONE) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
531 |
end |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
532 |
| _ => NONE); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
533 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
534 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
535 |
(* destruct nested patterns *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
536 |
|
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
537 |
fun encode_clause recur S T (pat, rhs) = |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
538 |
fold (fn x as (_, U) => fn t => |
57445 | 539 |
let val T = fastype_of t; |
69593 | 540 |
in Const (\<^const_name>\<open>case_abs\<close>, (U --> T) --> T) $ Term.absfree x t end) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
541 |
(Term.add_frees pat []) |
69593 | 542 |
(Const (\<^const_name>\<open>case_elem\<close>, S --> T --> S --> T) $ pat $ recur rhs); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
543 |
|
69593 | 544 |
fun encode_cases _ S T [] = Const (\<^const_name>\<open>case_nil\<close>, S --> T) |
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
545 |
| encode_cases recur S T (p :: ps) = |
69593 | 546 |
Const (\<^const_name>\<open>case_cons\<close>, (S --> T) --> (S --> T) --> S --> T) $ |
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
547 |
encode_clause recur S T p $ encode_cases recur S T ps; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
548 |
|
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
549 |
fun encode_case recur (t, ps as (pat, rhs) :: _) = |
52154 | 550 |
let |
551 |
val tT = fastype_of t; |
|
552 |
val T = fastype_of rhs; |
|
553 |
in |
|
69593 | 554 |
Const (\<^const_name>\<open>case_guard\<close>, \<^typ>\<open>bool\<close> --> tT --> (tT --> T) --> T) $ |
555 |
\<^term>\<open>True\<close> $ t $ (encode_cases recur (fastype_of pat) (fastype_of rhs) ps) |
|
52154 | 556 |
end |
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
557 |
| encode_case _ _ = case_error "encode_case"; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
558 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
559 |
fun strip_case' ctxt d (pat, rhs) = |
81506
f76a5849b570
tuned: more standard Name.build_context, although that is a bit longer;
wenzelm
parents:
81505
diff
changeset
|
560 |
(case dest_case ctxt d (Name.build_context (Term.declare_free_names pat)) rhs of |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
561 |
SOME (exp as Free _, clauses) => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
562 |
if Term.exists_subterm (curry (op aconv) exp) pat andalso |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
563 |
not (exists (fn (_, rhs') => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
564 |
Term.exists_subterm (curry (op aconv) exp) rhs') clauses) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
565 |
then |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
566 |
maps (strip_case' ctxt d) (map (fn (pat', rhs') => |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
567 |
(subst_free [(exp, pat')] pat, rhs')) clauses) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
568 |
else [(pat, rhs)] |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
569 |
| _ => [(pat, rhs)]); |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
570 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
571 |
fun strip_case ctxt d t = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
572 |
(case dest_case ctxt d Name.context t of |
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
573 |
SOME (x, clauses) => SOME (x, maps (strip_case' ctxt d) clauses) |
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
574 |
| NONE => NONE); |
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
575 |
|
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
576 |
fun strip_case_full ctxt d t = |
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
577 |
(case dest_case ctxt d Name.context t of |
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
578 |
SOME (x, clauses) => |
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
579 |
encode_case (strip_case_full ctxt d) |
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
580 |
(strip_case_full ctxt d x, maps (strip_case' ctxt d) clauses) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
581 |
| NONE => |
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
582 |
(case t of |
52154 | 583 |
t $ u => strip_case_full ctxt d t $ strip_case_full ctxt d u |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
71224
diff
changeset
|
584 |
| Abs _ => |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
71224
diff
changeset
|
585 |
let val (v, t') = Term.dest_abs_global t; |
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
71224
diff
changeset
|
586 |
in Term.absfree v (strip_case_full ctxt d t') end |
51677
d2b3372e6033
recur in the expression to be matched (do not rely on repetitive execution of a check phase);
traytel
parents:
51676
diff
changeset
|
587 |
| _ => t)); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
588 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
589 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
590 |
(* term uncheck *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
591 |
|
69593 | 592 |
val show_cases = Attrib.setup_config_bool \<^binding>\<open>show_cases\<close> (K true); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
593 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
594 |
fun uncheck_case ctxt ts = |
52285
da42b500a6aa
permissive uncheck -- allow printing of malformed terms (e.g. in error messages);
wenzelm
parents:
52265
diff
changeset
|
595 |
if Config.get ctxt show_cases |
da42b500a6aa
permissive uncheck -- allow printing of malformed terms (e.g. in error messages);
wenzelm
parents:
52265
diff
changeset
|
596 |
then map (fn t => if can Term.type_of t then strip_case_full ctxt true t else t) ts |
da42b500a6aa
permissive uncheck -- allow printing of malformed terms (e.g. in error messages);
wenzelm
parents:
52265
diff
changeset
|
597 |
else ts; |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
598 |
|
56362 | 599 |
val _ = Context.>> (Syntax_Phases.term_uncheck 1 "case" uncheck_case); |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
600 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
601 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
602 |
(* outer syntax commands *) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
603 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
604 |
fun print_case_translations ctxt = |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
605 |
let |
52155 | 606 |
val cases = map snd (Symtab.dest (cases_of ctxt)); |
607 |
val type_space = Proof_Context.type_space ctxt; |
|
608 |
||
609 |
val pretty_term = Syntax.pretty_term ctxt; |
|
610 |
||
611 |
fun pretty_data (data as (comb, ctrs)) = |
|
612 |
let |
|
613 |
val name = Tname_of_data data; |
|
614 |
val xname = Name_Space.extern ctxt type_space name; |
|
615 |
val markup = Name_Space.markup type_space name; |
|
616 |
val prt = |
|
617 |
(Pretty.block o Pretty.fbreaks) |
|
618 |
[Pretty.block [Pretty.mark_str (markup, xname), Pretty.str ":"], |
|
619 |
Pretty.block [Pretty.str "combinator:", Pretty.brk 1, pretty_term comb], |
|
620 |
Pretty.block (Pretty.str "constructors:" :: Pretty.brk 1 :: |
|
621 |
Pretty.commas (map pretty_term ctrs))]; |
|
622 |
in (xname, prt) end; |
|
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
623 |
in |
60924
610794dff23c
tuned signature, in accordance to sortBy in Scala;
wenzelm
parents:
59936
diff
changeset
|
624 |
Pretty.big_list "case translations:" (map #2 (sort_by #1 (map pretty_data cases))) |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
625 |
|> Pretty.writeln |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
626 |
end; |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
627 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
628 |
val _ = |
69593 | 629 |
Outer_Syntax.command \<^command_keyword>\<open>print_case_translations\<close> |
51673
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
630 |
"print registered case combinators and constructors" |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
631 |
(Scan.succeed (Toplevel.keep (print_case_translations o Toplevel.context_of))) |
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
632 |
|
4dfa00e264d8
separate data used for case translation from the datatype package
traytel
parents:
diff
changeset
|
633 |
end; |