Skip to content

Commit 0486959

Browse files
committed
WIP
1 parent 9b53bd3 commit 0486959

File tree

6 files changed

+217
-82
lines changed

6 files changed

+217
-82
lines changed

Defs.thy

Lines changed: 34 additions & 39 deletions
Original file line numberDiff line numberDiff line change
@@ -27,6 +27,13 @@ proof goal_cases
2727
qed (auto simp: eqvt_def isin_graph_aux_def)
2828
nominal_termination (eqvt) by lexicographic_order
2929

30+
nominal_function set_alts :: "alt_list \<Rightarrow> alt fset" where
31+
"set_alts ANil = {||}"
32+
| "set_alts (ACons alt alts) = finsert alt (set_alts alts)"
33+
apply (auto simp: eqvt_def set_alts_graph_aux_def)
34+
by (cases rule: term_alt_list_alt.exhaust(2)) auto
35+
nominal_termination (eqvt) by lexicographic_order
36+
3037
nominal_function head_ctor :: "term \<Rightarrow> bool" where
3138
"head_ctor (Var _) = False"
3239
| "head_ctor (Lam _ _ _) = False"
@@ -202,20 +209,20 @@ qed auto
202209
nominal_termination (eqvt) by lexicographic_order
203210

204211
abbreviation exhaustive :: "alt_list \<Rightarrow> \<Delta> \<Rightarrow> data_name \<Rightarrow> bool" where
205-
"exhaustive alts \<Delta> T \<equiv>
206-
(\<nexists>x e. MatchVar x e \<in> set_alt_list alts) \<longrightarrow>
207-
(\<forall>D \<tau> ks.
208-
(AxCtor D \<tau> \<in> set \<Delta> \<and> ctor_type \<tau> = Some (T, ks)) \<longrightarrow>
209-
(\<exists>tys vals e. MatchCtor D tys vals e \<in> set_alt_list alts)
210-
)"
212+
"exhaustive alts \<Delta> T \<equiv> \<forall>D \<tau> ks. (AxCtor D \<tau> \<in> set \<Delta> \<and> ctor_type \<tau> = Some (T, ks)) \<longrightarrow> (\<exists>tys vals e. MatchCtor D tys vals e |\<in>| set_alts alts)"
211213

212214
nominal_function zip_with :: "('a::pt \<Rightarrow> 'b::pt \<Rightarrow> 'c::pt) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where
213215
"zip_with _ [] _ = []"
214216
| "zip_with _ _ [] = []"
215217
| "zip_with f (a#as) (b#bs) = f a b # zip_with f as bs"
216218
proof goal_cases
217219
case (3 P x)
218-
then show ?case sorry
220+
then obtain f xs ys where P: "x = (f, xs, ys)" by (metis prod.exhaust)
221+
then show ?case using 3
222+
proof (cases xs)
223+
case (Cons a list)
224+
then show ?thesis using 3 P by (cases ys) auto
225+
qed simp
219226
qed (auto simp: eqvt_def zip_with_graph_aux_def)
220227
nominal_termination (eqvt) by lexicographic_order
221228

@@ -261,7 +268,6 @@ nominal_function (default "case_sum (\<lambda>x. Inl undefined) (case_sum (\<lam
261268
| "subst_alt_list ANil _ _ = ANil"
262269
| "subst_alt_list (ACons alt alts) e x = ACons (subst_alt alt e x) (subst_alt_list alts e x)"
263270

264-
| "atom y \<sharp> (e, x) \<Longrightarrow> subst_alt (MatchVar y t) e x = MatchVar y (subst_term t e x)"
265271
| "set (map atom tys @ map atom vals) \<sharp>* (e, x) \<Longrightarrow> subst_alt (MatchCtor D tys vals t) e x = MatchCtor D tys vals (subst_term t e x)"
266272
proof (goal_cases)
267273

@@ -308,31 +314,26 @@ proof (goal_cases)
308314
next
309315
case (Inr c)
310316
then obtain m e y where "c = (m, e, y)" by (metis prod.exhaust)
311-
then show ?thesis using 3(11,12) Inr outer fresh_star_insert
317+
then show ?thesis using 3 Inr outer fresh_star_insert
312318
apply (cases m rule: term_alt_list_alt.strong_exhaust(3)[of "(e, y)"])
313-
apply auto
314-
by blast
319+
by auto
315320
qed
316321
qed
317322
next
318-
case (54 y e x \<tau> e2 y' e' x' \<tau>' e2')
319-
have "(\<lambda> y : \<tau> . subst_term e2 e x) = (\<lambda> y' : \<tau>' . subst_term e2' e' x')" using Abs_sumC[OF 54(5,6) eqvt_at_term[OF 54(1)] eqvt_at_term[OF 54(2)]] 54(7) by fastforce
320-
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_def, symmetric, unfolded fun_eq_iff])
321-
next
322-
case (61 y e x k e2 y' e' x' k' e2')
323-
have "(\<Lambda> y : k . subst_term e2 e x) = (\<Lambda> y' : k' . subst_term e2' e' x')" using Abs_sumC[OF 61(5,6) eqvt_at_term[OF 61(1)] eqvt_at_term[OF 61(2)]] 61(7) by fastforce
323+
case (49 y e x \<tau> e2 y' e' x' \<tau>' e2')
324+
have "(\<lambda> y : \<tau> . subst_term e2 e x) = (\<lambda> y' : \<tau>' . subst_term e2' e' x')" using Abs_sumC[OF 49(5,6) eqvt_at_term[OF 49(1)] eqvt_at_term[OF 49(2)]] 49(7) by fastforce
324325
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_def, symmetric, unfolded fun_eq_iff])
325326
next
326-
case (67 y e x \<tau> e1 e2 y' e' x' \<tau>' e1' e2')
327-
have "Let y \<tau> (subst_term e1 e x) (subst_term e2 e x) = Let y' \<tau>' (subst_term e1' e' x') (subst_term e2' e' x')" using Abs_sumC[OF 67(9,10) eqvt_at_term[OF 67(2)] eqvt_at_term[OF 67(4)]] 67(11) by fastforce
327+
case (55 y e x k e2 y' e' x' k' e2')
328+
have "(\<Lambda> y : k . subst_term e2 e x) = (\<Lambda> y' : k' . subst_term e2' e' x')" using Abs_sumC[OF 55(5,6) eqvt_at_term[OF 55(1)] eqvt_at_term[OF 55(2)]] 55(7) by fastforce
328329
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_def, symmetric, unfolded fun_eq_iff])
329330
next
330-
case (79 y e x t y' e' x' t')
331-
have "MatchVar y (subst_term t e x) = MatchVar y' (subst_term t' e' x')" using Abs_sumC[OF 79(5,6) eqvt_at_term[OF 79(1)] eqvt_at_term[OF 79(2)]] 79(7) by fastforce
331+
case (60 y e x \<tau> e1 e2 y' e' x' \<tau>' e1' e2')
332+
have "Let y \<tau> (subst_term e1 e x) (subst_term e2 e x) = Let y' \<tau>' (subst_term e1' e' x') (subst_term e2' e' x')" using Abs_sumC[OF 60(9,10) eqvt_at_term[OF 60(2)] eqvt_at_term[OF 60(4)]] 60(11) by fastforce
332333
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_def, symmetric, unfolded fun_eq_iff])
333334
next
334-
case (81 tys vals e x D t tys' vals' e' x' D' t')
335-
have "MatchCtor D tys vals (subst_term t e x) = MatchCtor D' tys' vals' (subst_term t' e' x')" using Abs_sumC_star[OF 81(5,6) eqvt_at_term[OF 81(1)] eqvt_at_term[OF 81(2)]] 81(7) by fastforce
335+
case (69 tys vals e x D t tys' vals' e' x' D' t')
336+
have "MatchCtor D tys vals (subst_term t e x) = MatchCtor D' tys' vals' (subst_term t' e' x')" using Abs_sumC_star[OF 69(5,6) eqvt_at_term[OF 69(1)] eqvt_at_term[OF 69(2)]] 69(7) by fastforce
336337
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_def, symmetric, unfolded fun_eq_iff])
337338
} qed (auto simp: eqvt_def subst_term_subst_alt_list_subst_alt_graph_aux_def)
338339
nominal_termination (eqvt) by lexicographic_order
@@ -373,7 +374,6 @@ nominal_function (default "case_sum (\<lambda>x. Inl undefined) (case_sum (\<lam
373374
| "subst_alt_list_type ANil _ _ = ANil"
374375
| "subst_alt_list_type (ACons alt alts) \<tau> a = ACons (subst_alt_type alt \<tau> a) (subst_alt_list_type alts \<tau> a)"
375376

376-
| "atom y \<sharp> (\<tau>, a) \<Longrightarrow> subst_alt_type (MatchVar y e) \<tau> a = MatchVar y (subst_term_type e \<tau> a)"
377377
| "set (map atom tys @ map atom vals) \<sharp>* (\<tau>, a) \<Longrightarrow> subst_alt_type (MatchCtor D tys vals e) \<tau> a = MatchCtor D tys vals (subst_term_type e \<tau> a)"
378378
proof goal_cases
379379

@@ -420,31 +420,26 @@ proof goal_cases
420420
next
421421
case (Inr c)
422422
then obtain m e y where P: "c = (m, e, y)" by (metis prod.exhaust)
423-
then show ?thesis using Inr outer 3(11,12)
424-
apply (cases m rule: term_alt_list_alt.strong_exhaust(3)[of "(e, y)"]) apply simp
425-
using fresh_star_insert by blast
423+
then show ?thesis using Inr outer 3
424+
apply (cases m rule: term_alt_list_alt.strong_exhaust(3)[of "(e, y)"]) by simp
426425
qed
427426
qed
428427
next
429-
case (54 y \<tau> a \<tau>2 e2 y' \<tau>' a' \<tau>2' e2')
430-
have "(\<lambda> y : subst_type \<tau>2 \<tau> a . subst_term_type e2 \<tau> a) = (\<lambda> y' : subst_type \<tau>2' \<tau>' a' . subst_term_type e2' \<tau>' a')" using Abs_sumC[OF 54(5,6) eqvt_at_term[OF 54(1)] eqvt_at_term[OF 54(2)]] 54(7) by fastforce
428+
case (49 y \<tau> a \<tau>2 e2 y' \<tau>' a' \<tau>2' e2')
429+
have "(\<lambda> y : subst_type \<tau>2 \<tau> a . subst_term_type e2 \<tau> a) = (\<lambda> y' : subst_type \<tau>2' \<tau>' a' . subst_term_type e2' \<tau>' a')" using Abs_sumC[OF 49(5,6) eqvt_at_term[OF 49(1)] eqvt_at_term[OF 49(2)]] 49(7) by fastforce
431430
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_type_def, symmetric, unfolded fun_eq_iff])
432431
next
433-
case (61 b \<tau> a k e2 b' \<tau>' a' k' e2')
434-
have "(\<Lambda> b : k . subst_term_type e2 \<tau> a) = (\<Lambda> b' : k' . subst_term_type e2' \<tau>' a')" using Abs_sumC[OF 61(5,6) eqvt_at_term[OF 61(1)] eqvt_at_term[OF 61(2)]] 61(7) by fastforce
432+
case (55 b \<tau> a k e2 b' \<tau>' a' k' e2')
433+
have "(\<Lambda> b : k . subst_term_type e2 \<tau> a) = (\<Lambda> b' : k' . subst_term_type e2' \<tau>' a')" using Abs_sumC[OF 55(5,6) eqvt_at_term[OF 55(1)] eqvt_at_term[OF 55(2)]] 55(7) by fastforce
435434
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_type_def, symmetric, unfolded fun_eq_iff])
436435
next
437-
case (67 y \<tau> a \<tau>2 e1 e2 y' \<tau>' a' \<tau>2' e1' e2')
436+
case (60 y \<tau> a \<tau>2 e1 e2 y' \<tau>' a' \<tau>2' e1' e2')
438437
have "Let y (subst_type \<tau>2 \<tau> a) (subst_term_type e1 \<tau> a) (subst_term_type e2 \<tau> a) = Let y' (subst_type \<tau>2' \<tau>' a') (subst_term_type e1' \<tau>' a') (subst_term_type e2' \<tau>' a')"
439-
using Abs_sumC[OF 67(9,10) eqvt_at_term[OF 67(2)] eqvt_at_term[OF 67(4)]] 67(11) by fastforce
440-
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_type_def, symmetric, unfolded fun_eq_iff])
441-
next
442-
case (79 y \<tau> a e y' \<tau>' a' e')
443-
have "MatchVar y (subst_term_type e \<tau> a) = MatchVar y' (subst_term_type e' \<tau>' a')" using Abs_sumC[OF 79(5,6) eqvt_at_term[OF 79(1)] eqvt_at_term[OF 79(2)]] 79(7) by fastforce
438+
using Abs_sumC[OF 60(9,10) eqvt_at_term[OF 60(2)] eqvt_at_term[OF 60(4)]] 60(11) by fastforce
444439
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_type_def, symmetric, unfolded fun_eq_iff])
445440
next
446-
case (81 tys vals \<tau> a D e tys' vals' \<tau>' a' D' e')
447-
have "MatchCtor D tys vals (subst_term_type e \<tau> a) = MatchCtor D' tys' vals' (subst_term_type e' \<tau>' a')" using Abs_sumC_star[OF 81(5,6) eqvt_at_term[OF 81(1)] eqvt_at_term[OF 81(2)]] 81(7) by fastforce
441+
case (69 tys vals \<tau> a D e tys' vals' \<tau>' a' D' e')
442+
have "MatchCtor D tys vals (subst_term_type e \<tau> a) = MatchCtor D' tys' vals' (subst_term_type e' \<tau>' a')" using Abs_sumC_star[OF 69(5,6) eqvt_at_term[OF 69(1)] eqvt_at_term[OF 69(2)]] 69(7) by fastforce
448443
then show ?case by (auto simp: Abs_fresh_iff meta_eq_to_obj_eq[OF subst_term_type_def, symmetric, unfolded fun_eq_iff])
449444
} qed (auto simp: eqvt_def subst_term_type_subst_alt_list_type_subst_alt_type_graph_aux_def)
450445
nominal_termination (eqvt) by lexicographic_order

Defs2.thy

Lines changed: 79 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,79 @@
1+
theory Defs2
2+
imports Lemmas
3+
begin
4+
5+
nominal_function ctor_data_app_subst :: "\<tau> \<Rightarrow> bool" where
6+
"ctor_data_app_subst (TyVar _) = False"
7+
| "ctor_data_app_subst (TyData _) = True"
8+
| "ctor_data_app_subst TyArrow = False"
9+
| "ctor_data_app_subst (TyApp \<tau>1 _) = ctor_data_app_subst \<tau>1"
10+
| "ctor_data_app_subst (TyForall _ _ _) = False"
11+
proof goal_cases
12+
case (3 P x)
13+
then show ?case by (cases x rule: \<tau>.exhaust) auto
14+
qed (auto simp: eqvt_def ctor_data_app_subst_graph_aux_def)
15+
nominal_termination (eqvt) by lexicographic_order
16+
17+
nominal_function ctor_args :: "\<tau> \<Rightarrow> \<tau> list option" where
18+
"ctor_args (TyVar _) = None"
19+
| "ctor_args (TyData T) = Some []"
20+
| "ctor_args TyArrow = None"
21+
| "ctor_args (TyApp (TyApp TyArrow \<tau>1) \<tau>2) = (case ctor_args \<tau>2 of
22+
Some tys \<Rightarrow> Some (\<tau>1#tys)
23+
| None \<Rightarrow> None)"
24+
| "ctor_args (TyApp (TyApp (TyVar a) \<tau>1) \<tau>2) = (if ctor_data_app_subst (TyApp (TyApp (TyVar a) \<tau>1) \<tau>2) then Some [] else None)"
25+
| "ctor_args (TyApp (TyApp (TyData T) \<tau>1) \<tau>2) = (if ctor_data_app_subst (TyApp (TyApp (TyData T) \<tau>1) \<tau>2) then Some [] else None)"
26+
| "ctor_args (TyApp (TyApp (TyApp \<tau>1' \<tau>2') \<tau>1) \<tau>2) = (if ctor_data_app_subst (TyApp (TyApp (TyApp \<tau>1' \<tau>2') \<tau>1) \<tau>2) then Some [] else None)"
27+
| "ctor_args (TyApp (TyApp (TyForall a k e) \<tau>1) \<tau>2) = (if ctor_data_app_subst (TyApp (TyApp (TyForall a k e) \<tau>1) \<tau>2) then Some [] else None)"
28+
| "ctor_args (TyApp (TyVar a) \<tau>2) = (if ctor_data_app_subst (TyApp (TyVar a) \<tau>2) then Some [] else None)"
29+
| "ctor_args (TyApp (TyData T) \<tau>2) = (if ctor_data_app_subst (TyApp (TyData T) \<tau>2) then Some [] else None)"
30+
| "ctor_args (TyApp TyArrow \<tau>2) = (if ctor_data_app_subst (TyApp TyArrow \<tau>2) then Some [] else None)"
31+
| "ctor_args (TyApp (TyForall a k e) \<tau>2) = (if ctor_data_app_subst (TyApp (TyForall a k e) \<tau>2) then Some [] else None)"
32+
| "ctor_args (TyForall _ _ _) = None"
33+
proof goal_cases
34+
case 1
35+
then show ?case by (auto simp: eqvt_def ctor_args_graph_aux_def split!: option.splits)
36+
next
37+
case (3 P x)
38+
then show ?case
39+
proof (cases x rule: \<tau>.exhaust)
40+
case outer: (TyApp \<tau>1 \<tau>2)
41+
then show ?thesis using 3
42+
proof (cases \<tau>1 rule: \<tau>.exhaust)
43+
case (TyApp \<sigma>1 \<sigma>2)
44+
then show ?thesis using outer 3 by (cases \<sigma>1 rule: \<tau>.exhaust) auto
45+
qed auto
46+
qed auto
47+
qed auto
48+
nominal_termination (eqvt) by lexicographic_order
49+
50+
nominal_function subst_ctor :: "\<tau> \<Rightarrow> \<tau> list \<Rightarrow> \<tau> list option" where
51+
"subst_ctor (TyVar _) _ = None"
52+
| "subst_ctor TyArrow _ = None"
53+
| "subst_ctor (TyData _) [] = Some []"
54+
| "subst_ctor (TyData _) (_#_) = None"
55+
| "subst_ctor (TyApp \<tau>1 \<tau>2) [] = ctor_args (TyApp \<tau>1 \<tau>2)"
56+
| "subst_ctor (TyApp _ _) (_#_) = None"
57+
| "subst_ctor (TyForall _ _ _) [] = None"
58+
| "subst_ctor (TyForall a _ e) (\<tau>#\<tau>s) = subst_ctor e[\<tau>/a] \<tau>s"
59+
proof goal_cases
60+
case (3 P x)
61+
then obtain t xs where P: "x = (t, xs)" by fastforce
62+
then show ?case using 3
63+
proof (cases t rule: \<tau>.exhaust)
64+
case TyData
65+
then show ?thesis using P 3 by (cases xs) auto
66+
next
67+
case TyApp
68+
then show ?thesis using P 3 by (cases xs) auto
69+
next
70+
case TyForall
71+
then show ?thesis using P 3 by (cases xs) auto
72+
qed auto
73+
next
74+
case (39 a k e \<tau> \<tau>s a' k' e' \<tau>' \<tau>s')
75+
then show ?case by (metis Pair_inject \<tau>.eq_iff(5) list.inject subst_same(2))
76+
qed (auto simp: eqvt_def subst_ctor_graph_aux_def)
77+
nominal_termination (eqvt) by lexicographic_order
78+
79+
end

Lemmas.thy

Lines changed: 4 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -186,12 +186,9 @@ next
186186
case (8 y e x \<tau> e1 e2)
187187
then show ?case using fresh_PairD(2) fresh_at_base(2) by fastforce
188188
next
189-
case (11 y e x t)
190-
then show ?case using fresh_PairD(2) fresh_at_base(2) by fastforce
191-
next
192-
case (12 tys vals e x D t)
189+
case (11 tys vals e x D t)
193190
then have "atom x \<sharp> t" by (meson fresh_PairD(2) fresh_at_base(2) fresh_star_def term_alt_list_alt.fresh(11))
194-
then show ?case using 12 by simp
191+
then show ?case using 11 by simp
195192
qed auto
196193

197194
lemma fresh_subst_type_same: "atom a \<sharp> \<sigma> \<Longrightarrow> subst_type \<sigma> \<tau> a = \<sigma>"
@@ -209,9 +206,9 @@ proof (induction e \<tau> a and alts \<tau> a and alt \<tau> a rule: subst_term_
209206
case (7 b \<tau> a k e2)
210207
then show ?case by (simp add: "7.hyps" fresh_Pair fresh_at_base(2))
211208
next
212-
case (12 tys vals \<tau> a D e)
209+
case (11 tys vals \<tau> a D e)
213210
then have "atom a \<sharp> e" by (meson fresh_PairD(2) fresh_at_base(2) fresh_star_def term_alt_list_alt.fresh(11))
214-
then show ?case using 12 by auto
211+
then show ?case using 11 by auto
215212
qed (auto simp: fresh_subst_type_same)
216213

217214
lemma fresh_subst_context_same: "atom a \<sharp> \<Gamma> \<Longrightarrow> subst_context \<Gamma> \<tau> a = \<Gamma>"

Nominal2_Lemmas.thy

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -102,4 +102,9 @@ lemma Projr_permute: "\<exists>y. f = Inr y \<Longrightarrow> p \<bullet> projr
102102
lemma pair3_eqvt[simp]: "(a \<leftrightarrow> b) \<bullet> (x, y, z) = ((a \<leftrightarrow> b) \<bullet> x, (a \<leftrightarrow> b) \<bullet> y, (a \<leftrightarrow> b) \<bullet> z)"
103103
by (simp split: prod.splits)
104104

105+
lemma eqvt_fBall[eqvt]: "p \<bullet> fBall s f = fBall (p \<bullet> s) (p \<bullet> f)"
106+
apply auto
107+
apply (metis eqvt_bound eqvt_lambda fBallE in_fset_eqvt permute_pure)
108+
by (metis eqvt_apply fBallE fBallI in_fset_eqvt permute_pure)
109+
105110
end

Syntax.thy

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -36,7 +36,7 @@ and "alt_list" =
3636
| ACons "alt" "alt_list"
3737
and "alt" =
3838
MatchCtor "ctor_name" tys::"tyvar list" vals::"var list" e::"term" binds tys vals in e
39-
| MatchVar x::"var" e::"term" binds x in e
39+
(*| MatchVar x::"var" e::"term" binds x in e*)
4040

4141
nominal_datatype "binder" =
4242
BVar "var" "\<tau>"

0 commit comments

Comments
 (0)