Skip to content

Commit 8a6fd2d

Browse files
committed
Add basic axiomatization for BMV monads
1 parent be05f2b commit 8a6fd2d

File tree

3 files changed

+634
-15
lines changed

3 files changed

+634
-15
lines changed

Tools/bmv_monad_def.ML

Lines changed: 225 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,225 @@
1+
signature BMV_MONAD_DEF = sig
2+
type bmv_monad
3+
4+
type 'a bmv_monad_axioms = {
5+
Sb_Inj: 'a,
6+
Sb_comp_Injs: 'a list,
7+
Sb_comp: 'a,
8+
Sb_cong: 'a,
9+
Vrs_Injs: 'a list,
10+
Vrs_Sbs: 'a list
11+
};
12+
13+
type bmv_monad_model = {
14+
ops: typ list,
15+
bmv_ops: bmv_monad list,
16+
frees: typ list,
17+
lives: typ list,
18+
leader: int,
19+
Injs: (term list * (term * int) list) list,
20+
Sbs: term list,
21+
Maps: term option list,
22+
Vrs: term list list,
23+
tacs: (Proof.context -> tactic) bmv_monad_axioms list
24+
}
25+
26+
val ops_of_bmv_monad: bmv_monad -> typ list;
27+
val leader_of_bmv_monad: bmv_monad -> int;
28+
val frees_of_bmv_monad: bmv_monad -> typ list;
29+
val lives_of_bmv_monad: bmv_monad -> typ list;
30+
val Injs_of_bmv_monad: bmv_monad -> term list list;
31+
val Sbs_of_bmv_monad: bmv_monad -> term list;
32+
val Maps_of_bmv_monad: bmv_monad -> term option list;
33+
val Vrs_of_bmv_monad: bmv_monad -> term list list;
34+
35+
val morph_bmv_monad: morphism -> bmv_monad -> bmv_monad;
36+
37+
val bmv_monad_def: BNF_Def.inline_policy -> (Proof.context -> BNF_Def.fact_policy)
38+
-> (binding -> binding) -> bmv_monad_model -> local_theory -> bmv_monad
39+
end
40+
41+
structure BMV_Monad_Def : BMV_MONAD_DEF = struct
42+
43+
open MRBNF_Util
44+
45+
type 'a bmv_monad_axioms = {
46+
Sb_Inj: 'a,
47+
Sb_comp_Injs: 'a list,
48+
Sb_comp: 'a,
49+
Sb_cong: 'a,
50+
Vrs_Injs: 'a list,
51+
Vrs_Sbs: 'a list
52+
};
53+
54+
fun morph_bmv_axioms phi {
55+
Sb_Inj, Sb_comp_Injs, Sb_comp, Sb_cong, Vrs_Injs, Vrs_Sbs
56+
} = {
57+
Sb_Inj = Morphism.thm phi Sb_Inj,
58+
Sb_comp_Injs = map (Morphism.thm phi) Sb_comp_Injs,
59+
Sb_comp = Morphism.thm phi Sb_comp,
60+
Sb_cong = Morphism.thm phi Sb_cong,
61+
Vrs_Injs = map (Morphism.thm phi) Vrs_Injs,
62+
Vrs_Sbs = map (Morphism.thm phi) Vrs_Sbs
63+
} : thm bmv_monad_axioms
64+
65+
datatype bmv_monad = BMV of {
66+
ops: typ list,
67+
leader: int,
68+
frees: typ list,
69+
lives: typ list,
70+
Injs: term list list,
71+
Sbs: term list,
72+
Maps: term option list,
73+
Vrs: term list list,
74+
axioms: thm bmv_monad_axioms list
75+
}
76+
77+
fun Rep_bmv (BMV x) = x
78+
79+
val ops_of_bmv_monad = #ops o Rep_bmv
80+
val leader_of_bmv_monad = #leader o Rep_bmv
81+
val frees_of_bmv_monad = #frees o Rep_bmv
82+
val lives_of_bmv_monad = #lives o Rep_bmv
83+
val Injs_of_bmv_monad = #Injs o Rep_bmv
84+
val Sbs_of_bmv_monad = #Sbs o Rep_bmv
85+
val Maps_of_bmv_monad = #Maps o Rep_bmv
86+
val Vrs_of_bmv_monad = #Vrs o Rep_bmv
87+
88+
type bmv_monad_model = {
89+
ops: typ list,
90+
frees: typ list,
91+
lives: typ list,
92+
bmv_ops: bmv_monad list,
93+
leader: int,
94+
Injs: (term list * (term * int) list) list,
95+
Sbs: term list,
96+
Maps: term option list,
97+
Vrs: term list list,
98+
tacs: (Proof.context -> tactic) bmv_monad_axioms list
99+
}
100+
101+
fun morph_bmv_monad phi (BMV {
102+
ops, leader, frees, lives, Injs, Sbs, Maps, Vrs, axioms
103+
}) = BMV {
104+
ops = map (Morphism.typ phi) ops,
105+
leader = leader,
106+
frees = map (Morphism.typ phi) frees,
107+
lives = map (Morphism.typ phi) lives,
108+
Injs = map (map (Morphism.term phi)) Injs,
109+
Sbs = map (Morphism.term phi) Sbs,
110+
Maps = map (Option.map (Morphism.term phi)) Maps,
111+
Vrs = map (map (Morphism.term phi)) Vrs,
112+
axioms = map (morph_bmv_axioms phi) axioms
113+
}
114+
115+
fun prove_axioms (model: bmv_monad_model) lthy =
116+
let
117+
val Ts = #ops model @ maps ops_of_bmv_monad (#bmv_ops model);
118+
val Sbs = #Sbs model @ maps Sbs_of_bmv_monad (#bmv_ops model);
119+
val Injss = #Injs model @ map (rpair []) (maps Injs_of_bmv_monad (#bmv_ops model));
120+
val Vrss = #Vrs model @ maps Vrs_of_bmv_monad (#bmv_ops model);
121+
122+
val axioms = @{map 5} (fn T => fn (own_Injs, other_Injs) => fn Sb => fn Vrs => fn tacs =>
123+
let
124+
val (other_Injs, other_idxs) = split_list other_Injs;
125+
val Injs = own_Injs @ other_Injs;
126+
val ((((rhos, rhos'), aa), x), _) = lthy
127+
|> mk_Frees "\<rho>" (map fastype_of Injs)
128+
||>> mk_Frees "\<rho>'" (map fastype_of Injs)
129+
||>> mk_Frees "a" (map (fst o dest_funT o fastype_of) Injs)
130+
||>> apfst hd o mk_Frees "x" [T];
131+
val nown = length own_Injs;
132+
val (own_rhos, other_rhos) = chop nown rhos;
133+
val (own_rhos', other_rhos') = chop nown rhos';
134+
135+
val Sb_Inj = Goal.prove_sorry lthy [] [] (
136+
mk_Trueprop_eq (Term.list_comb (Sb, Injs), HOLogic.id_const T)
137+
) (fn {context=ctxt, ...} => #Sb_Inj tacs ctxt);
138+
139+
fun mk_small_prems rhos = map2 (fn rho => fn Inj => HOLogic.mk_Trueprop (mk_ordLess
140+
(mk_card_of (HOLogic.mk_Collect ("a", fst (dest_funT (fastype_of Inj)),
141+
HOLogic.mk_not (HOLogic.mk_eq (rho $ Bound 0, Inj $ Bound 0))
142+
)))
143+
(mk_card_of (HOLogic.mk_UNIV (fst (dest_funT (fastype_of Inj)))))
144+
)) rhos Injs;
145+
val small_prems = mk_small_prems rhos;
146+
val small_prems' = mk_small_prems rhos';
147+
148+
val Sb_comp_Injs = @{map 3} (fn Inj => fn rho => fn tac => Goal.prove_sorry lthy [] [] (
149+
fold_rev Logic.all rhos (fold_rev (curry Logic.mk_implies) small_prems (mk_Trueprop_eq (
150+
HOLogic.mk_comp (Term.list_comb (Sb, rhos), Inj), rho
151+
)))
152+
) (fn {context=ctxt, ...} => tac ctxt)) own_Injs own_rhos (#Sb_comp_Injs tacs);
153+
154+
val Sb_comp = Goal.prove_sorry lthy [] [] (fold_rev Logic.all (rhos @ rhos') (
155+
fold_rev (curry Logic.mk_implies) (small_prems @ small_prems') (mk_Trueprop_eq (
156+
HOLogic.mk_comp (Term.list_comb (Sb, rhos'), Term.list_comb (Sb, rhos)),
157+
Term.list_comb (Sb, map (fn rho => HOLogic.mk_comp (
158+
Term.list_comb (Sb, rhos'), rho
159+
)) own_rhos @ @{map 3} (fn rho => fn Sb => fn Injs =>
160+
HOLogic.mk_comp (Term.list_comb (Sb, map (fn Inj =>
161+
case List.find (fn rho' => fastype_of rho' = fastype_of Inj) rhos' of
162+
NONE => Inj | SOME t => t
163+
) (fst Injs @ map fst (snd Injs))), rho)
164+
) other_rhos (map (nth Sbs) other_idxs) (map (nth Injss) other_idxs))
165+
))
166+
)) (fn {context=ctxt, ...} => #Sb_comp tacs ctxt);
167+
168+
val Vrs_Injs = @{map 4} (fn Vrs => fn Inj => fn a => fn tac => Goal.prove_sorry lthy [] [] (
169+
Logic.all a (mk_Trueprop_eq (Vrs $ (Inj $ a), mk_singleton a))
170+
) (fn {context=ctxt, ...} => tac ctxt)) (take nown Vrs) own_Injs (take nown aa) (#Vrs_Injs tacs);
171+
172+
val Vrs_Sbs = map2 (fn Vr => fn tac => Goal.prove_sorry lthy [] [] (fold_rev Logic.all (rhos @ [x]) (
173+
fold_rev (curry Logic.mk_implies) small_prems (mk_Trueprop_eq (
174+
Vr $ (Term.list_comb (Sb, rhos) $ x),
175+
foldl1 mk_Un (@{map_filter 2} (fn rho => Option.map (fn Vrs' => mk_UNION (Vr $ x) (
176+
Term.abs ("a", HOLogic.dest_setT (snd (dest_funT (fastype_of Vrs')))) (
177+
Vrs' $ (rho $ Bound 0)
178+
)
179+
))) rhos (map SOME (take nown Vrs) @ map (fn idx =>
180+
List.find (fn t => body_type (fastype_of t) = body_type (fastype_of Vr)) (nth Vrss idx)
181+
) other_idxs))
182+
))
183+
)) (fn {context=ctxt, ...} => tac ctxt)) Vrs (#Vrs_Sbs tacs);
184+
185+
val Sb_cong = Goal.prove_sorry lthy [] [] (fold_rev Logic.all (rhos @ rhos' @ [x]) (
186+
fold_rev (curry Logic.mk_implies) (small_prems @ small_prems' @ @{map 4} (fn rho => fn rho' => fn Vrs => fn a =>
187+
Logic.all a (Logic.mk_implies (
188+
HOLogic.mk_Trueprop (HOLogic.mk_mem (a, Vrs $ x)),
189+
mk_Trueprop_eq (rho $ a, rho' $ a)
190+
))
191+
) rhos rhos' Vrs aa) (mk_Trueprop_eq (
192+
Term.list_comb (Sb, rhos) $ x,
193+
Term.list_comb (Sb, rhos') $ x
194+
)
195+
))) (fn {context=ctxt, ...} => #Sb_cong tacs ctxt);
196+
197+
in {
198+
Sb_Inj = Sb_Inj,
199+
Sb_comp_Injs = Sb_comp_Injs,
200+
Sb_comp = Sb_comp,
201+
Vrs_Injs = Vrs_Injs,
202+
Vrs_Sbs = Vrs_Sbs,
203+
Sb_cong = Sb_cong
204+
} end
205+
) (#ops model) (#Injs model) (#Sbs model) (#Vrs model) (#tacs model);
206+
in axioms end;
207+
208+
fun bmv_monad_def inline_policy fact_policy qualify (model: bmv_monad_model) lthy =
209+
let
210+
val axioms = prove_axioms model lthy;
211+
212+
val bmv = BMV {
213+
ops = #ops model @ maps (#ops o Rep_bmv) (#bmv_ops model),
214+
leader = #leader model,
215+
frees = #frees model,
216+
lives = #lives model,
217+
Injs = map (fn (xs, ys) => xs @ map fst ys) (#Injs model) @ maps (#Injs o Rep_bmv) (#bmv_ops model),
218+
Sbs = #Sbs model @ maps (#Sbs o Rep_bmv) (#bmv_ops model),
219+
Vrs = #Vrs model @ maps (#Vrs o Rep_bmv) (#bmv_ops model),
220+
Maps = #Maps model @ maps (#Maps o Rep_bmv) (#bmv_ops model),
221+
axioms = axioms @ maps (#axioms o Rep_bmv) (#bmv_ops model)
222+
} : bmv_monad;
223+
in bmv end
224+
225+
end

Tools/mrbnf_util.ML

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -11,6 +11,8 @@ sig
1111
val short_type_name: string -> string
1212
val swap: 'a * 'b -> 'b * 'a
1313

14+
val subst_typ_morphism: (typ * typ) list -> morphism
15+
1416
val mk_supp: term -> term
1517
val mk_supp_bound: term -> term
1618
val mk_imsupp: term -> term
@@ -88,6 +90,12 @@ fun mk_ex (x, T) t = HOLogic.mk_exists (x, T, t);
8890
fun mk_insert x S =
8991
Const (@{const_name Set.insert}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S;
9092

93+
fun subst_typ_morphism subst = Morphism.morphism "subst_typ" {
94+
binding = [], fact = [],
95+
term = [K (Term.subst_atomic_types subst)],
96+
typ = [K (Term.typ_subst_atomic subst)]
97+
};
98+
9199
fun mk_def_t_syn syn public b qualify name n rhs lthy =
92100
let
93101
val b' = qualify (Binding.name name);

0 commit comments

Comments
 (0)