Skip to content

Commit 1a1d851

Browse files
authored
Add BMV monads (#60)
2 parents ad5f48b + 13b4335 commit 1a1d851

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

47 files changed

+11341
-3893
lines changed

ROOT

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -22,6 +22,8 @@ session Binders in "thys" = Prelim +
2222
MRBNF_Composition
2323
MRBNF_FP
2424
MRBNF_Recursor
25+
Swapping
26+
Support
2527

2628
session Operations in "operations" = Untyped_Lambda_Calculus +
2729
theories
@@ -36,6 +38,8 @@ session Operations in "operations" = Untyped_Lambda_Calculus +
3638
VVSubst_Corec
3739
TVSubst
3840
Sugar
41+
BMV_Monad
42+
BMV_Fixpoint
3943

4044
session Tests in "tests" = Case_Studies +
4145
sessions
@@ -116,3 +120,7 @@ session System_Fsub in "case_studies/POPLmark" = Case_Studies +
116120
Pattern
117121
POPLmark_1B
118122
POPLmark_2B
123+
124+
session STLC in "case_studies/STLC" = Binders +
125+
theories
126+
STLC

Tools/binder_induction.ML

Lines changed: 9 additions & 13 deletions
Original file line numberDiff line numberDiff line change
@@ -121,14 +121,6 @@ fun gen_binder_context_tactic mod_cases simp def_insts arbitrary avoiding taking
121121
| NONE => error "Automatic detection of induction rule is not supported yet, please specify with rule:"
122122
);
123123

124-
fun dest_ordLess t =
125-
let val t' = case HOLogic.dest_mem t of
126-
(t', Const (@{const_name ordLess}, _)) => t'
127-
| _ => raise TERM ("dest_ordLess", [t])
128-
in HOLogic.dest_prod t' end
129-
fun dest_card_of (Const (@{const_name card_of}, _) $ t) = t
130-
| dest_card_of t = raise TERM ("dest_card_of", [t])
131-
132124
fun strip_all (Const (@{const_name HOL.All}, _) $ t) = (case t of
133125
Abs (x, T, t) => let val a = Free (x, T)
134126
in apfst (curry (op::) a) (strip_all (Term.subst_bound (a, t))) end
@@ -153,8 +145,8 @@ fun gen_binder_context_tactic mod_cases simp def_insts arbitrary avoiding taking
153145
val vars_prems = map (try (
154146
HOLogic.dest_Trueprop
155147
#> snd o strip_all
156-
#> fst o dest_ordLess
157-
#> dest_card_of
148+
#> fst o MRBNF_Util.dest_ordLess
149+
#> MRBNF_Util.dest_card_of
158150
#> fst o Term.strip_comb
159151
#> snd o Term.dest_Var
160152
#> HOLogic.dest_setT o range_type
@@ -463,6 +455,10 @@ fun gen_binder_context_tactic mod_cases simp def_insts arbitrary avoiding taking
463455
rotate_tac ~1
464456
],
465457
SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms de_Morgan_disj singleton_iff Int_Un_distrib Un_empty}),
458+
REPEAT_DETERM o CHANGED o EVERY' [
459+
SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms conj_assoc}),
460+
SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms conj_assoc[symmetric]})
461+
],
466462
REPEAT_DETERM o EVERY' [
467463
etac ctxt conjE,
468464
rotate_tac ~2
@@ -471,9 +467,9 @@ fun gen_binder_context_tactic mod_cases simp def_insts arbitrary avoiding taking
471467
Subgoal.FOCUS_PREMS (fn {context=ctxt, prems=inner_prems, ...} => EVERY1 [
472468
DETERM o resolve_tac ctxt (IH :: repeat_mp IH),
473469
IF_UNSOLVED o EVERY' [
474-
SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms Un_iff de_Morgan_disj singleton_iff Int_Un_distrib Un_empty}),
470+
K (Local_Defs.unfold0_tac ctxt @{thms notin_Un Un_iff de_Morgan_disj singleton_iff Int_Un_distrib Un_empty}),
475471
EVERY' (map (fn inner_prem => EVERY' (map (fn tac => DETERM o tac) [
476-
TRY o rtac ctxt conjI,
472+
REPEAT_DETERM o rtac ctxt conjI,
477473
let val x = length (fst (Logic.strip_horn (Thm.prop_of inner_prem)));
478474
in REPEAT_DETERM_N x o rtac ctxt @{thm induct_impliesI} end,
479475
REPEAT_DETERM o rtac ctxt @{thm induct_forallI},
@@ -484,7 +480,7 @@ fun gen_binder_context_tactic mod_cases simp def_insts arbitrary avoiding taking
484480
val simpset = Simplifier.add_cong @{thm imp_cong} (
485481
(BNF_Util.ss_only @{thms
486482
split_paired_All HOL.simp_thms all_simps HOL.induct_forall_def prod.inject induct_implies_equal_eq
487-
disjoint_single trans[OF arg_cong2[OF Int_commute refl, of "(=)"] disjoint_single]
483+
disjoint_single trans[OF arg_cong2[OF Int_commute refl, of "(=)"] disjoint_single] conj_assoc
488484
} ctxt)
489485
addsimprocs [@{simproc HOL.defined_All}, @{simproc HOL.defined_Ex}]
490486
)

Tools/binder_inductive.ML

Lines changed: 6 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -66,6 +66,7 @@ fun binder_inductive_cmd (((options, pred_name), binds_opt: (string * string lis
6666
(MRBNF_Util.subst_typ_morphism (snd (dest_Type (MRBNF_Def.T_of_mrbnf mrbnf)) ~~ Ts))
6767
mrbnf))
6868
| collect_params _ = [];
69+
6970
val pot_bind_Ts = foldl1 (fn (xs, ys) => if
7071
subset (op=) (xs, ys) then ys else union (op=) ys xs
7172
) (map collect_params param_Ts);
@@ -78,7 +79,7 @@ fun binder_inductive_cmd (((options, pred_name), binds_opt: (string * string lis
7879
| collect_binders (Abs (_, _, t)) = collect_binders t
7980
| collect_binders (t as (t1 $ t2)) = case try (dest_Type o Term.body_type o fastype_of) t of
8081
NONE => map2 (curry (op@)) (collect_binders t1) (collect_binders t2)
81-
| SOME (s, _) => (case MRBNF_Sugar.binder_sugar_of no_defs_lthy s of
82+
| SOME (s, _) => (case Binder_Sugar.binder_sugar_of no_defs_lthy s of
8283
NONE => map2 (curry (op@)) (collect_binders t1) (collect_binders t2)
8384
| SOME sugar =>
8485
let
@@ -115,7 +116,8 @@ fun binder_inductive_cmd (((options, pred_name), binds_opt: (string * string lis
115116
val intro = unvarify_global (Term.subst_TVars subst (Thm.prop_of intro));
116117
val binderss = map (distinct (op=)) (collect_binders intro);
117118
val xs_binders = map (fn binders => fold_rev (fn t => fn (a, b) => case t of
118-
Const (@{const_name insert}, _) $ x $ Const (@{const_name bot}, _) => (x::a, b)
119+
Const (@{const_name insert}, _) $ x $ Const (@{const_name bot}, _) => (case x of
120+
Bound _ => (a, b) | _ => (x::a, b))
119121
| _ => (a, t::b)
120122
) binders ([], [])) binderss;
121123
val binders = map (fn (xs, binders) =>
@@ -539,7 +541,7 @@ fun binder_inductive_cmd (((options, pred_name), binds_opt: (string * string lis
539541
map_filter (try dest_Type o snd) (fold Term.add_frees (flat bind_ts) [])
540542
);
541543

542-
val bset_bounds = maps (fn Type (s, _) => (case MRBNF_Sugar.binder_sugar_of lthy s of
544+
val bset_bounds = maps (fn Type (s, _) => (case Binder_Sugar.binder_sugar_of lthy s of
543545
SOME sugar => #bset_bounds sugar
544546
| NONE => [])
545547
| _ => []) param_Ts;
@@ -550,7 +552,7 @@ fun binder_inductive_cmd (((options, pred_name), binds_opt: (string * string lis
550552
REPEAT_DETERM o eresolve_tac ctxt [exE, conjE],
551553
hyp_subst_tac ctxt,
552554
REPEAT_DETERM o resolve_tac ctxt (
553-
@{thms conjI emp_bound iffD2[OF insert_bound] ordLeq_refl}
555+
@{thms conjI emp_bound iffD2[OF insert_bound_UNIV] ordLeq_refl}
554556
@ infinite_UNIVs @ [Un_bound, UN_bound]
555557
@ maps (fn thm => [thm, @{thm ordLess_ordLeq_trans} OF [thm]]) (
556558
maps set_bd_UNIVs_of_mr_bnfs binder_mr_bnfs

Tools/binder_sugar.ML

Lines changed: 88 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,88 @@
1+
signature BINDER_SUGAR = sig
2+
3+
type binder_sugar = {
4+
map_simps: thm list,
5+
set_simpss: thm list list,
6+
permute_simps: thm list,
7+
map_permute: thm option,
8+
subst_simps: thm list option,
9+
IImsupp_permute_commutes: thm list option,
10+
IImsupp_Diffs: thm list option,
11+
tvsubst_permute: thm option,
12+
bsetss: term option list list,
13+
bset_bounds: thm list,
14+
pset_ctors: thm list,
15+
mrbnf: MRBNF_Def.mrbnf,
16+
strong_induct: thm option,
17+
distinct: thm list,
18+
inject: thm list,
19+
ctors: (term * thm) list
20+
};
21+
22+
val morph_binder_sugar: morphism -> binder_sugar -> binder_sugar;
23+
24+
val binder_sugar_of: local_theory -> string -> binder_sugar option
25+
val register_binder_sugar: string -> binder_sugar -> local_theory -> local_theory
26+
27+
end
28+
29+
structure Binder_Sugar : BINDER_SUGAR = struct
30+
31+
type binder_sugar = {
32+
map_simps: thm list,
33+
set_simpss: thm list list,
34+
permute_simps: thm list,
35+
map_permute: thm option,
36+
subst_simps: thm list option,
37+
IImsupp_permute_commutes: thm list option,
38+
IImsupp_Diffs: thm list option,
39+
tvsubst_permute: thm option,
40+
bsetss: term option list list,
41+
bset_bounds: thm list,
42+
pset_ctors: thm list,
43+
mrbnf: MRBNF_Def.mrbnf,
44+
strong_induct: thm option,
45+
distinct: thm list,
46+
inject: thm list,
47+
ctors: (term * thm) list
48+
};
49+
50+
fun morph_binder_sugar phi { map_simps, permute_simps, map_permute, set_simpss, subst_simps, mrbnf,
51+
strong_induct, distinct, inject, ctors, bsetss, bset_bounds, IImsupp_permute_commutes, IImsupp_Diffs,
52+
tvsubst_permute, pset_ctors
53+
} = {
54+
map_simps = map (Morphism.thm phi) map_simps,
55+
permute_simps = map (Morphism.thm phi) permute_simps,
56+
map_permute = Option.map (Morphism.thm phi) map_permute,
57+
set_simpss = map (map (Morphism.thm phi)) set_simpss,
58+
subst_simps = Option.map (map (Morphism.thm phi)) subst_simps,
59+
IImsupp_permute_commutes = Option.map (map (Morphism.thm phi)) IImsupp_permute_commutes,
60+
IImsupp_Diffs = Option.map (map (Morphism.thm phi)) IImsupp_Diffs,
61+
tvsubst_permute = Option.map (Morphism.thm phi) tvsubst_permute,
62+
bsetss = map (map (Option.map (Morphism.term phi))) bsetss,
63+
bset_bounds = map (Morphism.thm phi) bset_bounds,
64+
pset_ctors = map (Morphism.thm phi) pset_ctors,
65+
mrbnf = MRBNF_Def.morph_mrbnf phi mrbnf,
66+
strong_induct = Option.map (Morphism.thm phi) strong_induct,
67+
distinct = map (Morphism.thm phi) distinct,
68+
inject = map (Morphism.thm phi) inject,
69+
ctors = map (MRBNF_Util.map_prod (Morphism.term phi) (Morphism.thm phi)) ctors
70+
} : binder_sugar;
71+
72+
structure Data = Generic_Data (
73+
type T = binder_sugar Symtab.table;
74+
val empty = Symtab.empty;
75+
fun merge data : T = Symtab.merge (K true) data;
76+
);
77+
78+
fun register_binder_sugar name sugar =
79+
Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.none}
80+
(fn phi => Data.map (Symtab.update (name, morph_binder_sugar phi sugar)));
81+
82+
fun binder_sugar_of_generic context =
83+
Option.map (morph_binder_sugar (Morphism.transfer_morphism (Context.theory_of context)))
84+
o Symtab.lookup (Data.get context);
85+
86+
val binder_sugar_of = binder_sugar_of_generic o Context.Proof;
87+
88+
end

0 commit comments

Comments
 (0)