author | Christian Urban <christian dot urban at kcl dot ac dot uk> |
Tue, 19 Feb 2013 05:38:46 +0000 | |
branch | Nominal2-Isabelle2013 |
changeset 3206 | fb201e383f1b |
parent 3199 | 93e7c1d8cc5c |
child 3214 | 13ab4f0a0b0e |
permissions | -rw-r--r-- |
2568
8193bbaa07fe
merged Nominal-General directory into Nominal; renamed Abs.thy to Nominal2_Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2559
diff
changeset
|
1 |
theory Nominal2_Abs |
8193bbaa07fe
merged Nominal-General directory into Nominal; renamed Abs.thy to Nominal2_Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2559
diff
changeset
|
2 |
imports "Nominal2_Base" |
2635
64b4cb2c2bf8
simple cases for string rule inductions
Christian Urban <urbanc@in.tum.de>
parents:
2616
diff
changeset
|
3 |
"~~/src/HOL/Quotient" |
64b4cb2c2bf8
simple cases for string rule inductions
Christian Urban <urbanc@in.tum.de>
parents:
2616
diff
changeset
|
4 |
"~~/src/HOL/Library/Quotient_List" |
64b4cb2c2bf8
simple cases for string rule inductions
Christian Urban <urbanc@in.tum.de>
parents:
2616
diff
changeset
|
5 |
"~~/src/HOL/Library/Quotient_Product" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
begin |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
|
2473
a3711f07449b
added the definition supp_rel (support w.r.t. a relation)
Christian Urban <urbanc@in.tum.de>
parents:
2471
diff
changeset
|
8 |
|
a3711f07449b
added the definition supp_rel (support w.r.t. a relation)
Christian Urban <urbanc@in.tum.de>
parents:
2471
diff
changeset
|
9 |
section {* Abstractions *} |
a3711f07449b
added the definition supp_rel (support w.r.t. a relation)
Christian Urban <urbanc@in.tum.de>
parents:
2471
diff
changeset
|
10 |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
fun |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
12 |
alpha_set |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
where |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
14 |
alpha_set[simp del]: |
3199 | 15 |
"alpha_set (bs, x) R f p (cs, y) \<longleftrightarrow> |
1465
4de35639fef0
added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1460
diff
changeset
|
16 |
f x - bs = f y - cs \<and> |
3199 | 17 |
(f x - bs) \<sharp>* p \<and> |
18 |
R (p \<bullet> x) y \<and> |
|
19 |
p \<bullet> bs = cs" |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
|
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
21 |
fun |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
22 |
alpha_res |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
23 |
where |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
24 |
alpha_res[simp del]: |
3199 | 25 |
"alpha_res (bs, x) R f p (cs, y) \<longleftrightarrow> |
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
26 |
f x - bs = f y - cs \<and> |
3199 | 27 |
(f x - bs) \<sharp>* p \<and> |
28 |
R (p \<bullet> x) y" |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
|
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
30 |
fun |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
31 |
alpha_lst |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
32 |
where |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
33 |
alpha_lst[simp del]: |
3199 | 34 |
"alpha_lst (bs, x) R f p (cs, y) \<longleftrightarrow> |
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
35 |
f x - set bs = f y - set cs \<and> |
3199 | 36 |
(f x - set bs) \<sharp>* p \<and> |
37 |
R (p \<bullet> x) y \<and> |
|
38 |
p \<bullet> bs = cs" |
|
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
39 |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
40 |
lemmas alphas = alpha_set.simps alpha_res.simps alpha_lst.simps |
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
41 |
|
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
42 |
notation |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
43 |
alpha_set ("_ \<approx>set _ _ _ _" [100, 100, 100, 100, 100] 100) and |
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
44 |
alpha_res ("_ \<approx>res _ _ _ _" [100, 100, 100, 100, 100] 100) and |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
45 |
alpha_lst ("_ \<approx>lst _ _ _ _" [100, 100, 100, 100, 100] 100) |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
46 |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
47 |
section {* Mono *} |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
48 |
|
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
49 |
lemma [mono]: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
50 |
shows "R1 \<le> R2 \<Longrightarrow> alpha_set bs R1 \<le> alpha_set bs R2" |
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
51 |
and "R1 \<le> R2 \<Longrightarrow> alpha_res bs R1 \<le> alpha_res bs R2" |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
52 |
and "R1 \<le> R2 \<Longrightarrow> alpha_lst cs R1 \<le> alpha_lst cs R2" |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
53 |
by (case_tac [!] bs, case_tac [!] cs) |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
54 |
(auto simp add: le_fun_def le_bool_def alphas) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
56 |
section {* Equivariance *} |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
57 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
58 |
lemma alpha_eqvt[eqvt]: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
59 |
shows "(bs, x) \<approx>set R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>set (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)" |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
60 |
and "(bs, x) \<approx>res R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>res (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)" |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
61 |
and "(ds, x) \<approx>lst R f q (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>lst (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> es, p \<bullet> y)" |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
62 |
unfolding alphas |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
63 |
unfolding permute_eqvt[symmetric] |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
64 |
unfolding set_eqvt[symmetric] |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
65 |
unfolding permute_fun_app_eq[symmetric] |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
66 |
unfolding Diff_eqvt[symmetric] |
3004 | 67 |
unfolding eq_eqvt[symmetric] |
68 |
unfolding fresh_star_eqvt[symmetric] |
|
3199 | 69 |
by (auto simp only: permute_bool_def) |
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
70 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
71 |
section {* Equivalence *} |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
72 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
73 |
lemma alpha_refl: |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
74 |
assumes a: "R x x" |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
75 |
shows "(bs, x) \<approx>set R f 0 (bs, x)" |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
76 |
and "(bs, x) \<approx>res R f 0 (bs, x)" |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
77 |
and "(cs, x) \<approx>lst R f 0 (cs, x)" |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
78 |
using a |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
79 |
unfolding alphas |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
80 |
unfolding fresh_star_def |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
81 |
by (simp_all add: fresh_zero_perm) |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
82 |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
83 |
lemma alpha_sym: |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
84 |
assumes a: "R (p \<bullet> x) y \<Longrightarrow> R (- p \<bullet> y) x" |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
85 |
shows "(bs, x) \<approx>set R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>set R f (- p) (bs, x)" |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
86 |
and "(bs, x) \<approx>res R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>res R f (- p) (bs, x)" |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
87 |
and "(ds, x) \<approx>lst R f p (es, y) \<Longrightarrow> (es, y) \<approx>lst R f (- p) (ds, x)" |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
88 |
unfolding alphas fresh_star_def |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
89 |
using a |
3199 | 90 |
by (auto simp add: fresh_minus_perm) |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
91 |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
92 |
lemma alpha_trans: |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
93 |
assumes a: "\<lbrakk>R (p \<bullet> x) y; R (q \<bullet> y) z\<rbrakk> \<Longrightarrow> R ((q + p) \<bullet> x) z" |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
94 |
shows "\<lbrakk>(bs, x) \<approx>set R f p (cs, y); (cs, y) \<approx>set R f q (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>set R f (q + p) (ds, z)" |
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
95 |
and "\<lbrakk>(bs, x) \<approx>res R f p (cs, y); (cs, y) \<approx>res R f q (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>res R f (q + p) (ds, z)" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
96 |
and "\<lbrakk>(es, x) \<approx>lst R f p (gs, y); (gs, y) \<approx>lst R f q (hs, z)\<rbrakk> \<Longrightarrow> (es, x) \<approx>lst R f (q + p) (hs, z)" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
97 |
using a |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
98 |
unfolding alphas fresh_star_def |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
99 |
by (simp_all add: fresh_plus_perm) |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
100 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
101 |
lemma alpha_sym_eqvt: |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
102 |
assumes a: "R (p \<bullet> x) y \<Longrightarrow> R y (p \<bullet> x)" |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
103 |
and b: "p \<bullet> R = R" |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
104 |
shows "(bs, x) \<approx>set R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>set R f (- p) (bs, x)" |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
105 |
and "(bs, x) \<approx>res R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>res R f (- p) (bs, x)" |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
106 |
and "(ds, x) \<approx>lst R f p (es, y) \<Longrightarrow> (es, y) \<approx>lst R f (- p) (ds, x)" |
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
107 |
apply(auto intro!: alpha_sym) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
108 |
apply(drule_tac [!] a) |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
109 |
apply(rule_tac [!] p="p" in permute_boolE) |
3199 | 110 |
apply(simp_all add: b permute_self) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
111 |
done |
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
112 |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
113 |
lemma alpha_set_trans_eqvt: |
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
114 |
assumes b: "(cs, y) \<approx>set R f q (ds, z)" |
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
115 |
and a: "(bs, x) \<approx>set R f p (cs, y)" |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
116 |
and d: "q \<bullet> R = R" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
117 |
and c: "\<lbrakk>R (p \<bullet> x) y; R y (- q \<bullet> z)\<rbrakk> \<Longrightarrow> R (p \<bullet> x) (- q \<bullet> z)" |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
118 |
shows "(bs, x) \<approx>set R f (q + p) (ds, z)" |
3199 | 119 |
apply(rule alpha_trans(1)[OF _ a b]) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
120 |
apply(drule c) |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
121 |
apply(rule_tac p="q" in permute_boolE) |
3199 | 122 |
apply(simp add: d permute_self) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
123 |
apply(rotate_tac -1) |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
124 |
apply(drule_tac p="q" in permute_boolI) |
3199 | 125 |
apply(simp add: d permute_self permute_eqvt[symmetric]) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
126 |
done |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
127 |
|
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
128 |
lemma alpha_res_trans_eqvt: |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
129 |
assumes b: "(cs, y) \<approx>res R f q (ds, z)" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
130 |
and a: "(bs, x) \<approx>res R f p (cs, y)" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
131 |
and d: "q \<bullet> R = R" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
132 |
and c: "\<lbrakk>R (p \<bullet> x) y; R y (- q \<bullet> z)\<rbrakk> \<Longrightarrow> R (p \<bullet> x) (- q \<bullet> z)" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
133 |
shows "(bs, x) \<approx>res R f (q + p) (ds, z)" |
3199 | 134 |
apply(rule alpha_trans(2)[OF _ a b]) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
135 |
apply(drule c) |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
136 |
apply(rule_tac p="q" in permute_boolE) |
3199 | 137 |
apply(simp add: d permute_self) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
138 |
apply(rotate_tac -1) |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
139 |
apply(drule_tac p="q" in permute_boolI) |
3199 | 140 |
apply(simp add: d permute_self permute_eqvt[symmetric]) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
141 |
done |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
142 |
|
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
143 |
lemma alpha_lst_trans_eqvt: |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
144 |
assumes b: "(cs, y) \<approx>lst R f q (ds, z)" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
145 |
and a: "(bs, x) \<approx>lst R f p (cs, y)" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
146 |
and d: "q \<bullet> R = R" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
147 |
and c: "\<lbrakk>R (p \<bullet> x) y; R y (- q \<bullet> z)\<rbrakk> \<Longrightarrow> R (p \<bullet> x) (- q \<bullet> z)" |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
148 |
shows "(bs, x) \<approx>lst R f (q + p) (ds, z)" |
3199 | 149 |
apply(rule alpha_trans(3)[OF _ a b]) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
150 |
apply(drule c) |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
151 |
apply(rule_tac p="q" in permute_boolE) |
3199 | 152 |
apply(simp add: d permute_self) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
153 |
apply(rotate_tac -1) |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
154 |
apply(drule_tac p="q" in permute_boolI) |
3199 | 155 |
apply(simp add: d permute_self permute_eqvt[symmetric]) |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
156 |
done |
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
157 |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
158 |
lemmas alpha_trans_eqvt = alpha_set_trans_eqvt alpha_res_trans_eqvt alpha_lst_trans_eqvt |
2313
25d2cdf7d7e4
transitivity proofs done
Christian Urban <urbanc@in.tum.de>
parents:
2311
diff
changeset
|
159 |
|
2311
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
160 |
|
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
161 |
section {* General Abstractions *} |
4da5c5c29009
work on transitivity proof
Christian Urban <urbanc@in.tum.de>
parents:
2302
diff
changeset
|
162 |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
163 |
fun |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
164 |
alpha_abs_set |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
165 |
where |
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
166 |
[simp del]: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
167 |
"alpha_abs_set (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>set (op=) supp p (cs, y))" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
|
1657 | 169 |
fun |
170 |
alpha_abs_lst |
|
171 |
where |
|
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
172 |
[simp del]: |
1657 | 173 |
"alpha_abs_lst (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>lst (op=) supp p (cs, y))" |
174 |
||
175 |
fun |
|
176 |
alpha_abs_res |
|
177 |
where |
|
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
178 |
[simp del]: |
1657 | 179 |
"alpha_abs_res (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op=) supp p (cs, y))" |
180 |
||
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
notation |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
182 |
alpha_abs_set (infix "\<approx>abs'_set" 50) and |
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
183 |
alpha_abs_lst (infix "\<approx>abs'_lst" 50) and |
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
184 |
alpha_abs_res (infix "\<approx>abs'_res" 50) |
1657 | 185 |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
186 |
lemmas alphas_abs = alpha_abs_set.simps alpha_abs_res.simps alpha_abs_lst.simps |
1657 | 187 |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
188 |
|
1657 | 189 |
lemma alphas_abs_refl: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
190 |
shows "(bs, x) \<approx>abs_set (bs, x)" |
1657 | 191 |
and "(bs, x) \<approx>abs_res (bs, x)" |
192 |
and "(cs, x) \<approx>abs_lst (cs, x)" |
|
193 |
unfolding alphas_abs |
|
194 |
unfolding alphas |
|
195 |
unfolding fresh_star_def |
|
196 |
by (rule_tac [!] x="0" in exI) |
|
197 |
(simp_all add: fresh_zero_perm) |
|
198 |
||
199 |
lemma alphas_abs_sym: |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
200 |
shows "(bs, x) \<approx>abs_set (cs, y) \<Longrightarrow> (cs, y) \<approx>abs_set (bs, x)" |
1657 | 201 |
and "(bs, x) \<approx>abs_res (cs, y) \<Longrightarrow> (cs, y) \<approx>abs_res (bs, x)" |
202 |
and "(ds, x) \<approx>abs_lst (es, y) \<Longrightarrow> (es, y) \<approx>abs_lst (ds, x)" |
|
203 |
unfolding alphas_abs |
|
204 |
unfolding alphas |
|
205 |
unfolding fresh_star_def |
|
206 |
by (erule_tac [!] exE, rule_tac [!] x="-p" in exI) |
|
207 |
(auto simp add: fresh_minus_perm) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
208 |
|
1657 | 209 |
lemma alphas_abs_trans: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
210 |
shows "\<lbrakk>(bs, x) \<approx>abs_set (cs, y); (cs, y) \<approx>abs_set (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>abs_set (ds, z)" |
1657 | 211 |
and "\<lbrakk>(bs, x) \<approx>abs_res (cs, y); (cs, y) \<approx>abs_res (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>abs_res (ds, z)" |
212 |
and "\<lbrakk>(es, x) \<approx>abs_lst (gs, y); (gs, y) \<approx>abs_lst (hs, z)\<rbrakk> \<Longrightarrow> (es, x) \<approx>abs_lst (hs, z)" |
|
213 |
unfolding alphas_abs |
|
214 |
unfolding alphas |
|
215 |
unfolding fresh_star_def |
|
216 |
apply(erule_tac [!] exE, erule_tac [!] exE) |
|
217 |
apply(rule_tac [!] x="pa + p" in exI) |
|
218 |
by (simp_all add: fresh_plus_perm) |
|
219 |
||
220 |
lemma alphas_abs_eqvt: |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
221 |
shows "(bs, x) \<approx>abs_set (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>abs_set (p \<bullet> cs, p \<bullet> y)" |
1657 | 222 |
and "(bs, x) \<approx>abs_res (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>abs_res (p \<bullet> cs, p \<bullet> y)" |
223 |
and "(ds, x) \<approx>abs_lst (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>abs_lst (p \<bullet> es, p \<bullet> y)" |
|
224 |
unfolding alphas_abs |
|
225 |
unfolding alphas |
|
226 |
unfolding set_eqvt[symmetric] |
|
227 |
unfolding supp_eqvt[symmetric] |
|
228 |
unfolding Diff_eqvt[symmetric] |
|
229 |
apply(erule_tac [!] exE) |
|
230 |
apply(rule_tac [!] x="p \<bullet> pa" in exI) |
|
3183
313e6f2cdd89
added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
Christian Urban <urbanc@in.tum.de>
parents:
3172
diff
changeset
|
231 |
by (auto simp only: fresh_star_permute_iff permute_eqvt[symmetric]) |
1657 | 232 |
|
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
233 |
|
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
234 |
section {* Strengthening the equivalence *} |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
235 |
|
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
236 |
lemma disjoint_right_eq: |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
237 |
assumes a: "A \<union> B1 = A \<union> B2" |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
238 |
and b: "A \<inter> B1 = {}" "A \<inter> B2 = {}" |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
239 |
shows "B1 = B2" |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
240 |
using a b |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
241 |
by (metis Int_Un_distrib2 Int_absorb2 Int_commute Un_upper2) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
242 |
|
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
243 |
lemma supp_property_res: |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
244 |
assumes a: "(as, x) \<approx>res (op =) supp p (as', x')" |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
245 |
shows "p \<bullet> (supp x \<inter> as) = supp x' \<inter> as'" |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
246 |
proof - |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
247 |
from a have "(supp x - as) \<sharp>* p" by (auto simp only: alphas) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
248 |
then have *: "p \<bullet> (supp x - as) = (supp x - as)" |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
249 |
by (simp add: atom_set_perm_eq) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
250 |
have "(supp x' - as') \<union> (supp x' \<inter> as') = supp x'" by auto |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
251 |
also have "\<dots> = supp (p \<bullet> x)" using a by (simp add: alphas) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
252 |
also have "\<dots> = p \<bullet> (supp x)" by (simp add: supp_eqvt) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
253 |
also have "\<dots> = p \<bullet> ((supp x - as) \<union> (supp x \<inter> as))" by auto |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
254 |
also have "\<dots> = (p \<bullet> (supp x - as)) \<union> (p \<bullet> (supp x \<inter> as))" by (simp add: union_eqvt) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
255 |
also have "\<dots> = (supp x - as) \<union> (p \<bullet> (supp x \<inter> as))" using * by simp |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
256 |
also have "\<dots> = (supp x' - as') \<union> (p \<bullet> (supp x \<inter> as))" using a by (simp add: alphas) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
257 |
finally have "(supp x' - as') \<union> (supp x' \<inter> as') = (supp x' - as') \<union> (p \<bullet> (supp x \<inter> as))" . |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
258 |
moreover |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
259 |
have "(supp x' - as') \<inter> (supp x' \<inter> as') = {}" by auto |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
260 |
moreover |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
261 |
have "(supp x - as) \<inter> (supp x \<inter> as) = {}" by auto |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
262 |
then have "p \<bullet> ((supp x - as) \<inter> (supp x \<inter> as) = {})" by (simp add: permute_bool_def) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
263 |
then have "(p \<bullet> (supp x - as)) \<inter> (p \<bullet> (supp x \<inter> as)) = {}" by (perm_simp) (simp) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
264 |
then have "(supp x - as) \<inter> (p \<bullet> (supp x \<inter> as)) = {}" using * by simp |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
265 |
then have "(supp x' - as') \<inter> (p \<bullet> (supp x \<inter> as)) = {}" using a by (simp add: alphas) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
266 |
ultimately show "p \<bullet> (supp x \<inter> as) = supp x' \<inter> as'" |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
267 |
by (auto dest: disjoint_right_eq) |
2712
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
268 |
qed |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
269 |
|
2674
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
270 |
lemma alpha_abs_res_stronger1_aux: |
2671
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
271 |
assumes asm: "(as, x) \<approx>res (op =) supp p' (as', x')" |
2669
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
272 |
shows "\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> (supp x \<inter> as) \<union> (supp x' \<inter> as')" |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
273 |
proof - |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
274 |
from asm have 0: "(supp x - as) \<sharp>* p'" by (auto simp only: alphas) |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
275 |
then have #: "p' \<bullet> (supp x - as) = (supp x - as)" |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
276 |
by (simp add: atom_set_perm_eq) |
2673
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
277 |
obtain p where *: "\<forall>b \<in> supp x. p \<bullet> b = p' \<bullet> b" and **: "supp p \<subseteq> supp x \<union> p' \<bullet> supp x" |
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
278 |
using set_renaming_perm2 by blast |
2669
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
279 |
from * have a: "p \<bullet> x = p' \<bullet> x" using supp_perm_perm_eq by auto |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
280 |
from 0 have 1: "(supp x - as) \<sharp>* p" using * |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
281 |
by (auto simp add: fresh_star_def fresh_perm) |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
282 |
then have 2: "(supp x - as) \<inter> supp p = {}" |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
283 |
by (auto simp add: fresh_star_def fresh_def) |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
284 |
have b: "supp x = (supp x - as) \<union> (supp x \<inter> as)" by auto |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
285 |
have "supp p \<subseteq> supp x \<union> p' \<bullet> supp x" using ** by simp |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
286 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> (p' \<bullet> ((supp x - as) \<union> (supp x \<inter> as)))" |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
287 |
using b by simp |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
288 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> ((p' \<bullet> (supp x - as)) \<union> (p' \<bullet> (supp x \<inter> as)))" |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
289 |
by (simp add: union_eqvt) |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
290 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> (p' \<bullet> (supp x \<inter> as))" |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
291 |
using # by auto |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
292 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> (supp x' \<inter> as')" using asm |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
293 |
by (simp add: supp_property_res) |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
294 |
finally have "supp p \<subseteq> (supp x - as) \<union> (supp x \<inter> as) \<union> (supp x' \<inter> as')" . |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
295 |
then |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
296 |
have "supp p \<subseteq> (supp x \<inter> as) \<union> (supp x' \<inter> as')" using 2 by auto |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
297 |
moreover |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
298 |
have "(as, x) \<approx>res (op =) supp p (as', x')" using asm 1 a by (simp add: alphas) |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
299 |
ultimately |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
300 |
show "\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> (supp x \<inter> as) \<union> (supp x' \<inter> as')" by blast |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
301 |
qed |
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
Christian Urban <urbanc@in.tum.de>
parents:
2668
diff
changeset
|
302 |
|
2712
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
303 |
lemma alpha_abs_res_minimal: |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
304 |
assumes asm: "(as, x) \<approx>res (op =) supp p (as', x')" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
305 |
shows "(as \<inter> supp x, x) \<approx>res (op =) supp p (as' \<inter> supp x', x')" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
306 |
using asm unfolding alpha_res by (auto simp add: Diff_Int) |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
307 |
|
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
308 |
lemma alpha_abs_res_abs_set: |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
309 |
assumes asm: "(as, x) \<approx>res (op =) supp p (as', x')" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
310 |
shows "(as \<inter> supp x, x) \<approx>set (op =) supp p (as' \<inter> supp x', x')" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
311 |
proof - |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
312 |
have c: "p \<bullet> x = x'" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
313 |
using alpha_abs_res_minimal[OF asm] unfolding alpha_res by clarify |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
314 |
then have a: "supp x - as \<inter> supp x = supp (p \<bullet> x) - as' \<inter> supp (p \<bullet> x)" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
315 |
using alpha_abs_res_minimal[OF asm] by (simp add: alpha_res) |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
316 |
have b: "(supp x - as \<inter> supp x) \<sharp>* p" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
317 |
using alpha_abs_res_minimal[OF asm] unfolding alpha_res by clarify |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
318 |
have "p \<bullet> (as \<inter> supp x) = as' \<inter> supp (p \<bullet> x)" |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
319 |
by (metis Int_commute asm c supp_property_res) |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
320 |
then show ?thesis using a b c unfolding alpha_set by simp |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
321 |
qed |
c66d288b9fa0
alpha_res implies alpha_set :)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2683
diff
changeset
|
322 |
|
2713
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
323 |
lemma alpha_abs_set_abs_res: |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
324 |
assumes asm: "(as \<inter> supp x, x) \<approx>set (op =) supp p (as' \<inter> supp x', x')" |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
325 |
shows "(as, x) \<approx>res (op =) supp p (as', x')" |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
326 |
using asm unfolding alphas by (auto simp add: Diff_Int) |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
327 |
|
2674
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
328 |
lemma alpha_abs_res_stronger1: |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
329 |
assumes asm: "(as, x) \<approx>res (op =) supp p' (as', x')" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
330 |
shows "\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as'" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
331 |
using alpha_abs_res_stronger1_aux[OF asm] by auto |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
332 |
|
2671
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
333 |
lemma alpha_abs_set_stronger1: |
2673
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
334 |
assumes asm: "(as, x) \<approx>set (op =) supp p' (as', x')" |
2671
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
335 |
shows "\<exists>p. (as, x) \<approx>set (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as'" |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
336 |
proof - |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
337 |
from asm have 0: "(supp x - as) \<sharp>* p'" by (auto simp only: alphas) |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
338 |
then have #: "p' \<bullet> (supp x - as) = (supp x - as)" |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
339 |
by (simp add: atom_set_perm_eq) |
2673
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
340 |
obtain p where *: "\<forall>b \<in> (supp x \<union> as). p \<bullet> b = p' \<bullet> b" |
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
341 |
and **: "supp p \<subseteq> (supp x \<union> as) \<union> p' \<bullet> (supp x \<union> as)" |
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
342 |
using set_renaming_perm2 by blast |
2671
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
343 |
from * have "\<forall>b \<in> supp x. p \<bullet> b = p' \<bullet> b" by blast |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
344 |
then have a: "p \<bullet> x = p' \<bullet> x" using supp_perm_perm_eq by auto |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
345 |
from * have "\<forall>b \<in> as. p \<bullet> b = p' \<bullet> b" by blast |
2673
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
346 |
then have zb: "p \<bullet> as = p' \<bullet> as" |
3104
f7c4b8e6918b
updated to explicit set type constructor (post Isabelle 3rd January)
Christian Urban <urbanc@in.tum.de>
parents:
3060
diff
changeset
|
347 |
apply(auto simp add: permute_set_def) |
2673
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
348 |
apply(rule_tac x="xa" in exI) |
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
349 |
apply(simp) |
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
350 |
done |
2671
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
351 |
have zc: "p' \<bullet> as = as'" using asm by (simp add: alphas) |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
352 |
from 0 have 1: "(supp x - as) \<sharp>* p" using * |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
353 |
by (auto simp add: fresh_star_def fresh_perm) |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
354 |
then have 2: "(supp x - as) \<inter> supp p = {}" |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
355 |
by (auto simp add: fresh_star_def fresh_def) |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
356 |
have b: "supp x = (supp x - as) \<union> (supp x \<inter> as)" by auto |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
357 |
have "supp p \<subseteq> supp x \<union> as \<union> p' \<bullet> supp x \<union> p' \<bullet> as" using ** using union_eqvt by blast |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
358 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> (p' \<bullet> ((supp x - as) \<union> (supp x \<inter> as))) \<union> p' \<bullet> as" |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
359 |
using b by simp |
2673
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
360 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> |
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
361 |
((p' \<bullet> (supp x - as)) \<union> (p' \<bullet> (supp x \<inter> as))) \<union> p' \<bullet> as" by (simp add: union_eqvt) |
2671
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
362 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> (p' \<bullet> (supp x \<inter> as)) \<union> p' \<bullet> as" |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
363 |
using # by auto |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
364 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> p' \<bullet> ((supp x \<inter> as) \<union> as)" using union_eqvt |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
365 |
by auto |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
366 |
also have "\<dots> = (supp x - as) \<union> (supp x \<inter> as) \<union> as \<union> p' \<bullet> as" |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
367 |
by (metis Int_commute Un_commute sup_inf_absorb) |
2673
87ebc706df67
made alpha_abs_set_stronger1 stronger
Christian Urban <urbanc@in.tum.de>
parents:
2671
diff
changeset
|
368 |
also have "\<dots> = (supp x - as) \<union> as \<union> p' \<bullet> as" by blast |
2671
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
369 |
finally have "supp p \<subseteq> (supp x - as) \<union> as \<union> p' \<bullet> as" . |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
370 |
then have "supp p \<subseteq> as \<union> p' \<bullet> as" using 2 by blast |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
371 |
moreover |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
372 |
have "(as, x) \<approx>set (op =) supp p (as', x')" using asm 1 a zb by (simp add: alphas) |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
373 |
ultimately |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
374 |
show "\<exists>p. (as, x) \<approx>set (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as'" using zc by blast |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
375 |
qed |
eef49daac6c8
alpha_abs_set_stronger1
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2669
diff
changeset
|
376 |
|
2674
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
377 |
lemma alpha_abs_lst_stronger1: |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
378 |
assumes asm: "(as, x) \<approx>lst (op =) supp p' (as', x')" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
379 |
shows "\<exists>p. (as, x) \<approx>lst (op =) supp p (as', x') \<and> supp p \<subseteq> set as \<union> set as'" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
380 |
proof - |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
381 |
from asm have 0: "(supp x - set as) \<sharp>* p'" by (auto simp only: alphas) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
382 |
then have #: "p' \<bullet> (supp x - set as) = (supp x - set as)" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
383 |
by (simp add: atom_set_perm_eq) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
384 |
obtain p where *: "\<forall>b \<in> (supp x \<union> set as). p \<bullet> b = p' \<bullet> b" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
385 |
and **: "supp p \<subseteq> (supp x \<union> set as) \<union> p' \<bullet> (supp x \<union> set as)" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
386 |
using set_renaming_perm2 by blast |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
387 |
from * have "\<forall>b \<in> supp x. p \<bullet> b = p' \<bullet> b" by blast |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
388 |
then have a: "p \<bullet> x = p' \<bullet> x" using supp_perm_perm_eq by auto |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
389 |
from * have "\<forall>b \<in> set as. p \<bullet> b = p' \<bullet> b" by blast |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
390 |
then have zb: "p \<bullet> as = p' \<bullet> as" by (induct as) (auto) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
391 |
have zc: "p' \<bullet> set as = set as'" using asm by (simp add: alphas set_eqvt) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
392 |
from 0 have 1: "(supp x - set as) \<sharp>* p" using * |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
393 |
by (auto simp add: fresh_star_def fresh_perm) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
394 |
then have 2: "(supp x - set as) \<inter> supp p = {}" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
395 |
by (auto simp add: fresh_star_def fresh_def) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
396 |
have b: "supp x = (supp x - set as) \<union> (supp x \<inter> set as)" by auto |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
397 |
have "supp p \<subseteq> supp x \<union> set as \<union> p' \<bullet> supp x \<union> p' \<bullet> set as" using ** using union_eqvt by blast |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
398 |
also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
399 |
(p' \<bullet> ((supp x - set as) \<union> (supp x \<inter> set as))) \<union> p' \<bullet> set as" using b by simp |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
400 |
also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
401 |
((p' \<bullet> (supp x - set as)) \<union> (p' \<bullet> (supp x \<inter> set as))) \<union> p' \<bullet> set as" by (simp add: union_eqvt) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
402 |
also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
403 |
(p' \<bullet> (supp x \<inter> set as)) \<union> p' \<bullet> set as" using # by auto |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
404 |
also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> p' \<bullet> ((supp x \<inter> set as) \<union> set as)" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
405 |
using union_eqvt by auto |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
406 |
also have "\<dots> = (supp x - set as) \<union> (supp x \<inter> set as) \<union> set as \<union> p' \<bullet> set as" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
407 |
by (metis Int_commute Un_commute sup_inf_absorb) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
408 |
also have "\<dots> = (supp x - set as) \<union> set as \<union> p' \<bullet> set as" by blast |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
409 |
finally have "supp p \<subseteq> (supp x - set as) \<union> set as \<union> p' \<bullet> set as" . |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
410 |
then have "supp p \<subseteq> set as \<union> p' \<bullet> set as" using 2 by blast |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
411 |
moreover |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
412 |
have "(as, x) \<approx>lst (op =) supp p (as', x')" using asm 1 a zb by (simp add: alphas) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
413 |
ultimately |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
414 |
show "\<exists>p. (as, x) \<approx>lst (op =) supp p (as', x') \<and> supp p \<subseteq> set as \<union> set as'" using zc by blast |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
415 |
qed |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
416 |
|
2674
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
417 |
lemma alphas_abs_stronger: |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
418 |
shows "(as, x) \<approx>abs_set (as', x') \<longleftrightarrow> (\<exists>p. (as, x) \<approx>set (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as')" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
419 |
and "(as, x) \<approx>abs_res (as', x') \<longleftrightarrow> (\<exists>p. (as, x) \<approx>res (op =) supp p (as', x') \<and> supp p \<subseteq> as \<union> as')" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
420 |
and "(bs, x) \<approx>abs_lst (bs', x') \<longleftrightarrow> |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
421 |
(\<exists>p. (bs, x) \<approx>lst (op =) supp p (bs', x') \<and> supp p \<subseteq> set bs \<union> set bs')" |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
422 |
apply(rule iffI) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
423 |
apply(auto simp add: alphas_abs alpha_abs_set_stronger1)[1] |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
424 |
apply(auto simp add: alphas_abs)[1] |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
425 |
apply(rule iffI) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
426 |
apply(auto simp add: alphas_abs alpha_abs_res_stronger1)[1] |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
427 |
apply(auto simp add: alphas_abs)[1] |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
428 |
apply(rule iffI) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
429 |
apply(auto simp add: alphas_abs alpha_abs_lst_stronger1)[1] |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
430 |
apply(auto simp add: alphas_abs)[1] |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
431 |
done |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
432 |
|
3058 | 433 |
lemma alpha_res_alpha_set: |
434 |
"(bs, x) \<approx>res op = supp p (cs, y) \<longleftrightarrow> (bs \<inter> supp x, x) \<approx>set op = supp p (cs \<inter> supp y, y)" |
|
435 |
using alpha_abs_set_abs_res alpha_abs_res_abs_set by blast |
|
436 |
||
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
437 |
section {* Quotient types *} |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
438 |
|
3172
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
439 |
quotient_type |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
440 |
'a abs_set = "(atom set \<times> 'a::pt)" / "alpha_abs_set" |
3199 | 441 |
apply(rule equivpI) |
3172
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
442 |
unfolding reflp_def refl_on_def symp_def sym_def transp_def trans_def |
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
443 |
by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:) |
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
444 |
|
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
445 |
quotient_type |
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
446 |
'b abs_res = "(atom set \<times> 'b::pt)" / "alpha_abs_res" |
3199 | 447 |
apply(rule equivpI) |
3172
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
448 |
unfolding reflp_def refl_on_def symp_def sym_def transp_def trans_def |
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
449 |
by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:) |
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
450 |
|
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
451 |
quotient_type |
4cf3a4d36799
Added workaround for broken quotient_type in tip isabelle.
Cezary Kaliszyk <cezarykaliszyk@gmail.com>
parents:
3157
diff
changeset
|
452 |
'c abs_lst = "(atom list \<times> 'c::pt)" / "alpha_abs_lst" |
1657 | 453 |
apply(rule_tac [!] equivpI) |
2592
98236fbd8aa6
updated to Isabelle 2nd December
Christian Urban <urbanc@in.tum.de>
parents:
2591
diff
changeset
|
454 |
unfolding reflp_def refl_on_def symp_def sym_def transp_def trans_def |
1657 | 455 |
by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
456 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
457 |
quotient_definition |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
458 |
Abs_set ("[_]set. _" [60, 60] 60) |
1932 | 459 |
where |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
460 |
"Abs_set::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_set" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
461 |
is |
3152
da59c94bed7e
updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents:
3104
diff
changeset
|
462 |
"Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)" . |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
463 |
|
1657 | 464 |
quotient_definition |
1932 | 465 |
Abs_res ("[_]res. _" [60, 60] 60) |
466 |
where |
|
1657 | 467 |
"Abs_res::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_res" |
468 |
is |
|
3152
da59c94bed7e
updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents:
3104
diff
changeset
|
469 |
"Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)" . |
1657 | 470 |
|
471 |
quotient_definition |
|
1932 | 472 |
Abs_lst ("[_]lst. _" [60, 60] 60) |
473 |
where |
|
1657 | 474 |
"Abs_lst::atom list \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_lst" |
475 |
is |
|
3152
da59c94bed7e
updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents:
3104
diff
changeset
|
476 |
"Pair::atom list \<Rightarrow> ('a::pt) \<Rightarrow> (atom list \<times> 'a)" . |
1657 | 477 |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
478 |
lemma [quot_respect]: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
479 |
shows "(op= ===> op= ===> alpha_abs_set) Pair Pair" |
1657 | 480 |
and "(op= ===> op= ===> alpha_abs_res) Pair Pair" |
481 |
and "(op= ===> op= ===> alpha_abs_lst) Pair Pair" |
|
482 |
unfolding fun_rel_def |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
483 |
by (auto intro: alphas_abs_refl) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
484 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
485 |
lemma [quot_respect]: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
486 |
shows "(op= ===> alpha_abs_set ===> alpha_abs_set) permute permute" |
1657 | 487 |
and "(op= ===> alpha_abs_res ===> alpha_abs_res) permute permute" |
488 |
and "(op= ===> alpha_abs_lst ===> alpha_abs_lst) permute permute" |
|
489 |
unfolding fun_rel_def |
|
490 |
by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
491 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
492 |
lemma Abs_eq_iff: |
3058 | 493 |
shows "[bs]set. x = [bs']set. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>set (op =) supp p (bs', y))" |
494 |
and "[bs]res. x = [bs']res. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op =) supp p (bs', y))" |
|
495 |
and "[cs]lst. x = [cs']lst. y \<longleftrightarrow> (\<exists>p. (cs, x) \<approx>lst (op =) supp p (cs', y))" |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
496 |
by (lifting alphas_abs) |
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
497 |
|
2674
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
498 |
lemma Abs_eq_iff2: |
3058 | 499 |
shows "[bs]set. x = [bs']set. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>set (op=) supp p (bs', y) \<and> supp p \<subseteq> bs \<union> bs')" |
500 |
and "[bs]res. x = [bs']res. y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op=) supp p (bs', y) \<and> supp p \<subseteq> bs \<union> bs')" |
|
501 |
and "[cs]lst. x = [cs']lst. y \<longleftrightarrow> (\<exists>p. (cs, x) \<approx>lst (op=) supp p (cs', y) \<and> supp p \<subseteq> set cs \<union> set cs')" |
|
2674
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
502 |
by (lifting alphas_abs_stronger) |
3c79dfec1cf0
derived stronger Abs_eq_iff2 theorems
Christian Urban <urbanc@in.tum.de>
parents:
2673
diff
changeset
|
503 |
|
3024 | 504 |
|
2713
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
505 |
lemma Abs_eq_res_set: |
3058 | 506 |
shows "[bs]res. x = [cs]res. y \<longleftrightarrow> [bs \<inter> supp x]set. x = [cs \<inter> supp y]set. y" |
3024 | 507 |
unfolding Abs_eq_iff alpha_res_alpha_set by rule |
2713
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
508 |
|
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
509 |
lemma Abs_eq_res_supp: |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
510 |
assumes asm: "supp x \<subseteq> bs" |
3058 | 511 |
shows "[as]res. x = [as \<inter> bs]res. x" |
2713
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
512 |
unfolding Abs_eq_iff alphas |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
513 |
apply (rule_tac x="0::perm" in exI) |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
514 |
apply (simp add: fresh_star_zero) |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
515 |
using asm by blast |
a84999edbcb3
More properties that relate abs_res and abs_set. Also abs_res with less binders.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2712
diff
changeset
|
516 |
|
3199 | 517 |
lemma Abs_exhausts[cases type]: |
3058 | 518 |
shows "(\<And>as (x::'a::pt). y1 = [as]set. x \<Longrightarrow> P1) \<Longrightarrow> P1" |
519 |
and "(\<And>as (x::'a::pt). y2 = [as]res. x \<Longrightarrow> P2) \<Longrightarrow> P2" |
|
520 |
and "(\<And>bs (x::'a::pt). y3 = [bs]lst. x \<Longrightarrow> P3) \<Longrightarrow> P3" |
|
1686
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
521 |
by (lifting prod.exhaust[where 'a="atom set" and 'b="'a"] |
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
522 |
prod.exhaust[where 'a="atom set" and 'b="'a"] |
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
523 |
prod.exhaust[where 'a="atom list" and 'b="'a"]) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
524 |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
525 |
instantiation abs_set :: (pt) pt |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
526 |
begin |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
527 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
528 |
quotient_definition |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
529 |
"permute_abs_set::perm \<Rightarrow> ('a::pt abs_set) \<Rightarrow> 'a abs_set" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
530 |
is |
3152
da59c94bed7e
updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents:
3104
diff
changeset
|
531 |
"permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)" |
da59c94bed7e
updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents:
3104
diff
changeset
|
532 |
by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
533 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
534 |
lemma permute_Abs_set[simp]: |
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
535 |
fixes x::"'a::pt" |
3058 | 536 |
shows "(p \<bullet> ([as]set. x)) = [p \<bullet> as]set. (p \<bullet> x)" |
1657 | 537 |
by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"]) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
538 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
539 |
instance |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
540 |
apply(default) |
3199 | 541 |
apply(case_tac [!] x) |
1657 | 542 |
apply(simp_all) |
543 |
done |
|
544 |
||
545 |
end |
|
546 |
||
547 |
instantiation abs_res :: (pt) pt |
|
548 |
begin |
|
549 |
||
550 |
quotient_definition |
|
551 |
"permute_abs_res::perm \<Rightarrow> ('a::pt abs_res) \<Rightarrow> 'a abs_res" |
|
552 |
is |
|
553 |
"permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)" |
|
3152
da59c94bed7e
updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents:
3104
diff
changeset
|
554 |
by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt) |
1657 | 555 |
|
556 |
lemma permute_Abs_res[simp]: |
|
557 |
fixes x::"'a::pt" |
|
3058 | 558 |
shows "(p \<bullet> ([as]res. x)) = [p \<bullet> as]res. (p \<bullet> x)" |
1657 | 559 |
by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"]) |
560 |
||
561 |
instance |
|
562 |
apply(default) |
|
3199 | 563 |
apply(case_tac [!] x) |
1657 | 564 |
apply(simp_all) |
565 |
done |
|
566 |
||
567 |
end |
|
568 |
||
569 |
instantiation abs_lst :: (pt) pt |
|
570 |
begin |
|
571 |
||
572 |
quotient_definition |
|
573 |
"permute_abs_lst::perm \<Rightarrow> ('a::pt abs_lst) \<Rightarrow> 'a abs_lst" |
|
574 |
is |
|
575 |
"permute:: perm \<Rightarrow> (atom list \<times> 'a::pt) \<Rightarrow> (atom list \<times> 'a::pt)" |
|
3152
da59c94bed7e
updated to Isabelle version April 1
Christian Urban <urbanc@in.tum.de>
parents:
3104
diff
changeset
|
576 |
by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt) |
1657 | 577 |
|
578 |
lemma permute_Abs_lst[simp]: |
|
579 |
fixes x::"'a::pt" |
|
3058 | 580 |
shows "(p \<bullet> ([as]lst. x)) = [p \<bullet> as]lst. (p \<bullet> x)" |
1657 | 581 |
by (lifting permute_prod.simps[where 'a="atom list" and 'b="'a"]) |
582 |
||
583 |
instance |
|
584 |
apply(default) |
|
3199 | 585 |
apply(case_tac [!] x) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
586 |
apply(simp_all) |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
587 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
588 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
589 |
end |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
590 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
591 |
lemmas permute_Abs[eqvt] = permute_Abs_set permute_Abs_res permute_Abs_lst |
1657 | 592 |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
593 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
594 |
lemma Abs_swap1: |
1662 | 595 |
assumes a1: "a \<notin> (supp x) - bs" |
596 |
and a2: "b \<notin> (supp x) - bs" |
|
3058 | 597 |
shows "[bs]set. x = [(a \<rightleftharpoons> b) \<bullet> bs]set. ((a \<rightleftharpoons> b) \<bullet> x)" |
598 |
and "[bs]res. x = [(a \<rightleftharpoons> b) \<bullet> bs]res. ((a \<rightleftharpoons> b) \<bullet> x)" |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
599 |
unfolding Abs_eq_iff |
1662 | 600 |
unfolding alphas |
601 |
unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] |
|
602 |
unfolding fresh_star_def fresh_def |
|
603 |
unfolding swap_set_not_in[OF a1 a2] |
|
604 |
using a1 a2 |
|
605 |
by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI) |
|
606 |
(auto simp add: supp_perm swap_atom) |
|
607 |
||
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
608 |
lemma Abs_swap2: |
1662 | 609 |
assumes a1: "a \<notin> (supp x) - (set bs)" |
610 |
and a2: "b \<notin> (supp x) - (set bs)" |
|
3058 | 611 |
shows "[bs]lst. x = [(a \<rightleftharpoons> b) \<bullet> bs]lst. ((a \<rightleftharpoons> b) \<bullet> x)" |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
612 |
unfolding Abs_eq_iff |
1662 | 613 |
unfolding alphas |
614 |
unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] set_eqvt[symmetric] |
|
615 |
unfolding fresh_star_def fresh_def |
|
616 |
unfolding swap_set_not_in[OF a1 a2] |
|
617 |
using a1 a2 |
|
618 |
by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI) |
|
619 |
(auto simp add: supp_perm swap_atom) |
|
620 |
||
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
621 |
lemma Abs_supports: |
3058 | 622 |
shows "((supp x) - as) supports ([as]set. x)" |
623 |
and "((supp x) - as) supports ([as]res. x)" |
|
624 |
and "((supp x) - set bs) supports ([bs]lst. x)" |
|
1662 | 625 |
unfolding supports_def |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
626 |
unfolding permute_Abs |
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
627 |
by (simp_all add: Abs_swap1[symmetric] Abs_swap2[symmetric]) |
1657 | 628 |
|
1686
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
629 |
function |
3199 | 630 |
supp_set :: "('a::pt) abs_set \<Rightarrow> atom set" and |
631 |
supp_res :: "('a::pt) abs_res \<Rightarrow> atom set" and |
|
632 |
supp_lst :: "('a::pt) abs_lst \<Rightarrow> atom set" |
|
1686
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
633 |
where |
3058 | 634 |
"supp_set ([as]set. x) = supp x - as" |
3199 | 635 |
| "supp_res ([as]res. x) = supp x - as" |
636 |
| "supp_lst (Abs_lst cs x) = (supp x) - (set cs)" |
|
637 |
apply(simp_all add: Abs_eq_iff alphas_abs alphas) |
|
638 |
apply(case_tac x) |
|
639 |
apply(case_tac a) |
|
1686
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
640 |
apply(simp) |
3199 | 641 |
apply(case_tac b) |
642 |
apply(case_tac a) |
|
643 |
apply(simp) |
|
644 |
apply(case_tac ba) |
|
645 |
apply(simp) |
|
1686
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
646 |
done |
1657 | 647 |
|
3199 | 648 |
termination |
3058 | 649 |
by lexicographic_order |
1686
7b3dd407f6b3
got rid of the aux-function on the raw level, by defining it with function on the quotient level
Christian Urban <urbanc@in.tum.de>
parents:
1673
diff
changeset
|
650 |
|
2663
54aade5d0fe6
moved high level code from LamTest into the main libraries.
Christian Urban <urbanc@in.tum.de>
parents:
2659
diff
changeset
|
651 |
lemma supp_funs_eqvt[eqvt]: |
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
652 |
shows "(p \<bullet> supp_set x) = supp_set (p \<bullet> x)" |
1657 | 653 |
and "(p \<bullet> supp_res y) = supp_res (p \<bullet> y)" |
654 |
and "(p \<bullet> supp_lst z) = supp_lst (p \<bullet> z)" |
|
3199 | 655 |
apply(case_tac x) |
656 |
apply(simp) |
|
657 |
apply(case_tac y) |
|
658 |
apply(simp) |
|
659 |
apply(case_tac z) |
|
660 |
apply(simp) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
661 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
662 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
663 |
lemma Abs_fresh_aux: |
3058 | 664 |
shows "a \<sharp> [bs]set. x \<Longrightarrow> a \<sharp> supp_set ([bs]set. x)" |
665 |
and "a \<sharp> [bs]res. x \<Longrightarrow> a \<sharp> supp_res ([bs]res. x)" |
|
666 |
and "a \<sharp> [cs]lst. x \<Longrightarrow> a \<sharp> supp_lst ([cs]lst. x)" |
|
1932 | 667 |
by (rule_tac [!] fresh_fun_eqvt_app) |
2663
54aade5d0fe6
moved high level code from LamTest into the main libraries.
Christian Urban <urbanc@in.tum.de>
parents:
2659
diff
changeset
|
668 |
(auto simp only: eqvt_def eqvts_raw) |
1657 | 669 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
670 |
lemma Abs_supp_subset1: |
1657 | 671 |
assumes a: "finite (supp x)" |
3058 | 672 |
shows "(supp x) - as \<subseteq> supp ([as]set. x)" |
673 |
and "(supp x) - as \<subseteq> supp ([as]res. x)" |
|
674 |
and "(supp x) - (set bs) \<subseteq> supp ([bs]lst. x)" |
|
1657 | 675 |
unfolding supp_conv_fresh |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
676 |
by (auto dest!: Abs_fresh_aux) |
1932 | 677 |
(simp_all add: fresh_def supp_finite_atom_set a) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
678 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
679 |
lemma Abs_supp_subset2: |
1657 | 680 |
assumes a: "finite (supp x)" |
3058 | 681 |
shows "supp ([as]set. x) \<subseteq> (supp x) - as" |
682 |
and "supp ([as]res. x) \<subseteq> (supp x) - as" |
|
683 |
and "supp ([bs]lst. x) \<subseteq> (supp x) - (set bs)" |
|
1932 | 684 |
by (rule_tac [!] supp_is_subset) |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
685 |
(simp_all add: Abs_supports a) |
1478
1ea4ca823266
added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents:
1470
diff
changeset
|
686 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
687 |
lemma Abs_finite_supp: |
1657 | 688 |
assumes a: "finite (supp x)" |
3058 | 689 |
shows "supp ([as]set. x) = (supp x) - as" |
690 |
and "supp ([as]res. x) = (supp x) - as" |
|
691 |
and "supp ([bs]lst. x) = (supp x) - (set bs)" |
|
692 |
using Abs_supp_subset1[OF a] Abs_supp_subset2[OF a] |
|
693 |
by blast+ |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
694 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
695 |
lemma supp_Abs: |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
696 |
fixes x::"'a::fs" |
3058 | 697 |
shows "supp ([as]set. x) = (supp x) - as" |
698 |
and "supp ([as]res. x) = (supp x) - as" |
|
699 |
and "supp ([bs]lst. x) = (supp x) - (set bs)" |
|
700 |
by (simp_all add: Abs_finite_supp finite_supp) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
701 |
|
2469
4a6e78bd9de9
renamed alpha_gen -> alpha_set and Abs -> Abs_set etc
Christian Urban <urbanc@in.tum.de>
parents:
2468
diff
changeset
|
702 |
instance abs_set :: (fs) fs |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
703 |
apply(default) |
3199 | 704 |
apply(case_tac x) |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
705 |
apply(simp add: supp_Abs finite_supp) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
706 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
707 |
|
1657 | 708 |
instance abs_res :: (fs) fs |
709 |
apply(default) |
|
3199 | 710 |
apply(case_tac x) |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
711 |
apply(simp add: supp_Abs finite_supp) |
1657 | 712 |
done |
713 |
||
714 |
instance abs_lst :: (fs) fs |
|
715 |
apply(default) |
|
3199 | 716 |
apply(case_tac x) |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
717 |
apply(simp add: supp_Abs finite_supp) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
718 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
719 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
720 |
lemma Abs_fresh_iff: |
1657 | 721 |
fixes x::"'a::fs" |
3058 | 722 |
shows "a \<sharp> [bs]set. x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)" |
723 |
and "a \<sharp> [bs]res. x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)" |
|
724 |
and "a \<sharp> [cs]lst. x \<longleftrightarrow> a \<in> (set cs) \<or> (a \<notin> (set cs) \<and> a \<sharp> x)" |
|
1657 | 725 |
unfolding fresh_def |
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
726 |
unfolding supp_Abs |
1657 | 727 |
by auto |
1460
0fd03936dedb
merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1451
diff
changeset
|
728 |
|
2591
35c570891a3a
isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents:
2584
diff
changeset
|
729 |
lemma Abs_fresh_star_iff: |
35c570891a3a
isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents:
2584
diff
changeset
|
730 |
fixes x::"'a::fs" |
3058 | 731 |
shows "as \<sharp>* ([bs]set. x) \<longleftrightarrow> (as - bs) \<sharp>* x" |
732 |
and "as \<sharp>* ([bs]res. x) \<longleftrightarrow> (as - bs) \<sharp>* x" |
|
733 |
and "as \<sharp>* ([cs]lst. x) \<longleftrightarrow> (as - set cs) \<sharp>* x" |
|
2591
35c570891a3a
isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents:
2584
diff
changeset
|
734 |
unfolding fresh_star_def |
35c570891a3a
isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents:
2584
diff
changeset
|
735 |
by (auto simp add: Abs_fresh_iff) |
35c570891a3a
isarfied some of the high-level proofs
Christian Urban <urbanc@in.tum.de>
parents:
2584
diff
changeset
|
736 |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
737 |
lemma Abs_fresh_star: |
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
738 |
fixes x::"'a::fs" |
3058 | 739 |
shows "as \<subseteq> as' \<Longrightarrow> as \<sharp>* ([as']set. x)" |
740 |
and "as \<subseteq> as' \<Longrightarrow> as \<sharp>* ([as']res. x)" |
|
741 |
and "bs \<subseteq> set bs' \<Longrightarrow> bs \<sharp>* ([bs']lst. x)" |
|
2491
d0961e6d6881
more consistent naming in Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2489
diff
changeset
|
742 |
unfolding fresh_star_def |
2584
1eac050a36f4
completely different method fro deriving the exhaust lemma
Christian Urban <urbanc@in.tum.de>
parents:
2574
diff
changeset
|
743 |
by(auto simp add: Abs_fresh_iff) |
2468 | 744 |
|
2730
eebc24b9cf39
added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents:
2713
diff
changeset
|
745 |
lemma Abs_fresh_star2: |
eebc24b9cf39
added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents:
2713
diff
changeset
|
746 |
fixes x::"'a::fs" |
3058 | 747 |
shows "as \<inter> bs = {} \<Longrightarrow> as \<sharp>* ([bs]set. x) \<longleftrightarrow> as \<sharp>* x" |
748 |
and "as \<inter> bs = {} \<Longrightarrow> as \<sharp>* ([bs]res. x) \<longleftrightarrow> as \<sharp>* x" |
|
749 |
and "cs \<inter> set ds = {} \<Longrightarrow> cs \<sharp>* ([ds]lst. x) \<longleftrightarrow> cs \<sharp>* x" |
|
2730
eebc24b9cf39
added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents:
2713
diff
changeset
|
750 |
unfolding fresh_star_def Abs_fresh_iff |
eebc24b9cf39
added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents:
2713
diff
changeset
|
751 |
by auto |
eebc24b9cf39
added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents:
2713
diff
changeset
|
752 |
|
eebc24b9cf39
added a lemma about fresh_star and Abs
Christian Urban <urbanc@in.tum.de>
parents:
2713
diff
changeset
|
753 |
|
3058 | 754 |
section {* Abstractions of single atoms *} |
755 |
||
3199 | 756 |
|
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
757 |
lemma Abs1_eq: |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
758 |
fixes x y::"'a::fs" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
759 |
shows "[{atom a}]set. x = [{atom a}]set. y \<longleftrightarrow> x = y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
760 |
and "[{atom a}]res. x = [{atom a}]res. y \<longleftrightarrow> x = y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
761 |
and "[[atom a]]lst. x = [[atom a]]lst. y \<longleftrightarrow> x = y" |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
762 |
unfolding Abs_eq_iff2 alphas |
3199 | 763 |
by (auto simp add: supp_perm_singleton fresh_star_def fresh_zero_perm) |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
764 |
|
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
765 |
lemma Abs1_eq_iff_fresh: |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
766 |
fixes x y::"'a::fs" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
767 |
and a b c::"'b::at" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
768 |
assumes "atom c \<sharp> (a, b, x, y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
769 |
shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
770 |
and "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
771 |
and "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
772 |
proof - |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
773 |
have "[{atom a}]set. x = (a \<leftrightarrow> c) \<bullet> ([{atom a}]set. x)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
774 |
by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
775 |
then have "[{atom a}]set. x = [{atom c}]set. ((a \<leftrightarrow> c) \<bullet> x)" by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
776 |
moreover |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
777 |
have "[{atom b}]set. y = (b \<leftrightarrow> c) \<bullet> ([{atom b}]set. y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
778 |
by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
779 |
then have "[{atom b}]set. y = [{atom c}]set. ((b \<leftrightarrow> c) \<bullet> y)" by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
780 |
ultimately |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
781 |
show "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
782 |
by (simp add: Abs1_eq) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
783 |
next |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
784 |
have "[{atom a}]res. x = (a \<leftrightarrow> c) \<bullet> ([{atom a}]res. x)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
785 |
by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
786 |
then have "[{atom a}]res. x = [{atom c}]res. ((a \<leftrightarrow> c) \<bullet> x)" by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
787 |
moreover |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
788 |
have "[{atom b}]res. y = (b \<leftrightarrow> c) \<bullet> ([{atom b}]res. y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
789 |
by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
790 |
then have "[{atom b}]res. y = [{atom c}]res. ((b \<leftrightarrow> c) \<bullet> y)" by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
791 |
ultimately |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
792 |
show "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
793 |
by (simp add: Abs1_eq) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
794 |
next |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
795 |
have "[[atom a]]lst. x = (a \<leftrightarrow> c) \<bullet> ([[atom a]]lst. x)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
796 |
by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
797 |
then have "[[atom a]]lst. x = [[atom c]]lst. ((a \<leftrightarrow> c) \<bullet> x)" by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
798 |
moreover |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
799 |
have "[[atom b]]lst. y = (b \<leftrightarrow> c) \<bullet> ([[atom b]]lst. y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
800 |
by (rule_tac flip_fresh_fresh[symmetric]) (simp_all add: Abs_fresh_iff assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
801 |
then have "[[atom b]]lst. y = [[atom c]]lst. ((b \<leftrightarrow> c) \<bullet> y)" by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
802 |
ultimately |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
803 |
show "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
804 |
by (simp add: Abs1_eq) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
805 |
qed |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
806 |
|
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
807 |
lemma Abs1_eq_iff_all: |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
808 |
fixes x y::"'a::fs" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
809 |
and z::"'c::fs" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
810 |
and a b::"'b::at" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
811 |
shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (\<forall>c. atom c \<sharp> (a, b, x, y, z) \<longrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
812 |
and "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (\<forall>c. atom c \<sharp> (a, b, x, y, z) \<longrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
813 |
and "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (\<forall>c. atom c \<sharp> (a, b, x, y, z) \<longrightarrow> (a \<leftrightarrow> c) \<bullet> x = (b \<leftrightarrow> c) \<bullet> y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
814 |
apply(auto) |
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
815 |
apply(simp add: Abs1_eq_iff_fresh(1)[symmetric]) |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
816 |
apply(rule_tac ?'a="'b::at" and x="(a, b, x, y, z)" in obtain_fresh) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
817 |
apply(drule_tac x="aa" in spec) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
818 |
apply(simp) |
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
819 |
apply(subst Abs1_eq_iff_fresh(1)) |
3199 | 820 |
apply(auto simp add: fresh_Pair)[2] |
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
821 |
apply(simp add: Abs1_eq_iff_fresh(2)[symmetric]) |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
822 |
apply(rule_tac ?'a="'b::at" and x="(a, b, x, y, z)" in obtain_fresh) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
823 |
apply(drule_tac x="aa" in spec) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
824 |
apply(simp) |
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
825 |
apply(subst Abs1_eq_iff_fresh(2)) |
3199 | 826 |
apply(auto simp add: fresh_Pair)[2] |
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
827 |
apply(simp add: Abs1_eq_iff_fresh(3)[symmetric]) |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
828 |
apply(rule_tac ?'a="'b::at" and x="(a, b, x, y, z)" in obtain_fresh) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
829 |
apply(drule_tac x="aa" in spec) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
830 |
apply(simp) |
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
831 |
apply(subst Abs1_eq_iff_fresh(3)) |
3199 | 832 |
apply(auto simp add: fresh_Pair)[2] |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
833 |
done |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
834 |
|
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
835 |
lemma Abs1_eq_iff: |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
836 |
fixes x y::"'a::fs" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
837 |
and a b::"'b::at" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
838 |
shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
839 |
and "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
840 |
and "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)" |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
841 |
proof - |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
842 |
{ assume "a = b" |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
843 |
then have "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a = b \<and> x = y)" by (simp add: Abs1_eq) |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
844 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
845 |
moreover |
3199 | 846 |
{ assume *: "a \<noteq> b" and **: "[{atom a}]set. x = [{atom b}]set. y" |
847 |
have #: "atom a \<sharp> [{atom b}]set. y" by (simp add: **[symmetric] Abs_fresh_iff) |
|
848 |
have "[{atom a}]set. ((a \<leftrightarrow> b) \<bullet> y) = (a \<leftrightarrow> b) \<bullet> ([{atom b}]set. y)" by (simp) |
|
849 |
also have "\<dots> = [{atom b}]set. y" |
|
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
850 |
by (rule flip_fresh_fresh) (simp add: #, simp add: Abs_fresh_iff) |
3199 | 851 |
also have "\<dots> = [{atom a}]set. x" using ** by simp |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
852 |
finally have "a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" using # * by (simp add: Abs1_eq Abs_fresh_iff) |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
853 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
854 |
moreover |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
855 |
{ assume *: "a \<noteq> b" and **: "x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" |
3199 | 856 |
have "[{atom a}]set. x = [{atom a}]set. ((a \<leftrightarrow> b) \<bullet> y)" using ** by simp |
857 |
also have "\<dots> = (a \<leftrightarrow> b) \<bullet> ([{atom b}]set. y)" by (simp add: permute_set_def assms) |
|
858 |
also have "\<dots> = [{atom b}]set. y" |
|
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
859 |
by (rule flip_fresh_fresh) (simp add: Abs_fresh_iff **, simp add: Abs_fresh_iff) |
3199 | 860 |
finally have "[{atom a}]set. x = [{atom b}]set. y" . |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
861 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
862 |
ultimately |
3199 | 863 |
show "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)" |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
864 |
by blast |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
865 |
next |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
866 |
{ assume "a = b" |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
867 |
then have "Abs_res {atom a} x = Abs_res {atom b} y \<longleftrightarrow> (a = b \<and> x = y)" by (simp add: Abs1_eq) |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
868 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
869 |
moreover |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
870 |
{ assume *: "a \<noteq> b" and **: "Abs_res {atom a} x = Abs_res {atom b} y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
871 |
have #: "atom a \<sharp> Abs_res {atom b} y" by (simp add: **[symmetric] Abs_fresh_iff) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
872 |
have "Abs_res {atom a} ((a \<leftrightarrow> b) \<bullet> y) = (a \<leftrightarrow> b) \<bullet> (Abs_res {atom b} y)" by (simp add: assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
873 |
also have "\<dots> = Abs_res {atom b} y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
874 |
by (rule flip_fresh_fresh) (simp add: #, simp add: Abs_fresh_iff) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
875 |
also have "\<dots> = Abs_res {atom a} x" using ** by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
876 |
finally have "a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" using # * by (simp add: Abs1_eq Abs_fresh_iff) |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
877 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
878 |
moreover |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
879 |
{ assume *: "a \<noteq> b" and **: "x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
880 |
have "Abs_res {atom a} x = Abs_res {atom a} ((a \<leftrightarrow> b) \<bullet> y)" using ** by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
881 |
also have "\<dots> = (a \<leftrightarrow> b) \<bullet> Abs_res {atom b} y" by (simp add: permute_set_def assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
882 |
also have "\<dots> = Abs_res {atom b} y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
883 |
by (rule flip_fresh_fresh) (simp add: Abs_fresh_iff **, simp add: Abs_fresh_iff) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
884 |
finally have "Abs_res {atom a} x = Abs_res {atom b} y" . |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
885 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
886 |
ultimately |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
887 |
show "Abs_res {atom a} x = Abs_res {atom b} y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)" |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
888 |
by blast |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
889 |
next |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
890 |
{ assume "a = b" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
891 |
then have "Abs_lst [atom a] x = Abs_lst [atom b] y \<longleftrightarrow> (a = b \<and> x = y)" by (simp add: Abs1_eq) |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
892 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
893 |
moreover |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
894 |
{ assume *: "a \<noteq> b" and **: "Abs_lst [atom a] x = Abs_lst [atom b] y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
895 |
have #: "atom a \<sharp> Abs_lst [atom b] y" by (simp add: **[symmetric] Abs_fresh_iff) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
896 |
have "Abs_lst [atom a] ((a \<leftrightarrow> b) \<bullet> y) = (a \<leftrightarrow> b) \<bullet> (Abs_lst [atom b] y)" by (simp add: assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
897 |
also have "\<dots> = Abs_lst [atom b] y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
898 |
by (rule flip_fresh_fresh) (simp add: #, simp add: Abs_fresh_iff) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
899 |
also have "\<dots> = Abs_lst [atom a] x" using ** by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
900 |
finally have "a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" using # * by (simp add: Abs1_eq Abs_fresh_iff) |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
901 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
902 |
moreover |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
903 |
{ assume *: "a \<noteq> b" and **: "x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
904 |
have "Abs_lst [atom a] x = Abs_lst [atom a] ((a \<leftrightarrow> b) \<bullet> y)" using ** by simp |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
905 |
also have "\<dots> = (a \<leftrightarrow> b) \<bullet> Abs_lst [atom b] y" by (simp add: assms) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
906 |
also have "\<dots> = Abs_lst [atom b] y" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
907 |
by (rule flip_fresh_fresh) (simp add: Abs_fresh_iff **, simp add: Abs_fresh_iff) |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
908 |
finally have "Abs_lst [atom a] x = Abs_lst [atom b] y" . |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
909 |
} |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
910 |
ultimately |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
911 |
show "Abs_lst [atom a] x = Abs_lst [atom b] y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<leftrightarrow> b) \<bullet> y \<and> atom a \<sharp> y)" |
2679
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
912 |
by blast |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
913 |
qed |
e003e5e36bae
added Minimal file to test things
Christian Urban <urbanc@in.tum.de>
parents:
2674
diff
changeset
|
914 |
|
2683
42c0d011a177
ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents:
2679
diff
changeset
|
915 |
lemma Abs1_eq_iff': |
42c0d011a177
ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents:
2679
diff
changeset
|
916 |
fixes x::"'a::fs" |
3191
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
917 |
and a b::"'b::at" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
918 |
shows "[{atom a}]set. x = [{atom b}]set. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> (b \<leftrightarrow> a) \<bullet> x = y \<and> atom b \<sharp> x)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
919 |
and "[{atom a}]res. x = [{atom b}]res. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> (b \<leftrightarrow> a) \<bullet> x = y \<and> atom b \<sharp> x)" |
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
Christian Urban <urbanc@in.tum.de>
parents:
3183
diff
changeset
|
920 |
and "[[atom a]]lst. x = [[atom b]]lst. y \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> (b \<leftrightarrow> a) \<bullet> x = y \<and> atom b \<sharp> x)" |
2683
42c0d011a177
ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents:
2679
diff
changeset
|
921 |
using assms by (auto simp add: Abs1_eq_iff fresh_permute_left) |
42c0d011a177
ported some of the old proofs to serve as testcases
Christian Urban <urbanc@in.tum.de>
parents:
2679
diff
changeset
|
922 |
|
2468 | 923 |
|
3192
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
924 |
ML {* |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
925 |
fun alpha_single_simproc thm _ ss ctrm = |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
926 |
let |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
927 |
val ctxt = Simplifier.the_context ss |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
928 |
val thy = Proof_Context.theory_of ctxt |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
929 |
val _ $ (_ $ x) $ (_ $ y) = term_of ctrm |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
930 |
val cvrs = union (op =) (Term.add_frees x []) (Term.add_frees y []) |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
931 |
|> filter (fn (_, ty) => Sign.of_sort thy (ty, @{sort fs})) |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
932 |
|> map Free |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
933 |
|> HOLogic.mk_tuple |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
934 |
|> Thm.cterm_of thy |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
935 |
val cvrs_ty = ctyp_of_term cvrs |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
936 |
val thm' = thm |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
937 |
|> Drule.instantiate' [NONE, NONE, SOME cvrs_ty] [NONE, NONE, NONE, NONE, SOME cvrs] |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
938 |
in |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
939 |
SOME thm' |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
940 |
end |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
941 |
*} |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
942 |
|
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
943 |
simproc_setup alpha_set ("[{atom a}]set. x = [{atom b}]set. y") = |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
944 |
{* alpha_single_simproc @{thm Abs1_eq_iff_all(1)[THEN eq_reflection]} *} |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
945 |
|
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
946 |
simproc_setup alpha_res ("[{atom a}]res. x = [{atom b}]res. y") = |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
947 |
{* alpha_single_simproc @{thm Abs1_eq_iff_all(2)[THEN eq_reflection]} *} |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
948 |
|
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
949 |
simproc_setup alpha_lst ("[[atom a]]lst. x = [[atom b]]lst. y") = |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
950 |
{* alpha_single_simproc @{thm Abs1_eq_iff_all(3)[THEN eq_reflection]} *} |
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
951 |
|
14c7d7e29c44
added a simproc for alpha-equivalence to the simplifier
Christian Urban <urbanc@in.tum.de>
parents:
3191
diff
changeset
|
952 |
|
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
953 |
subsection {* Renaming of bodies of abstractions *} |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
954 |
|
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
955 |
lemma Abs_rename_set: |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
956 |
fixes x::"'a::fs" |
2659
619ecb57db38
strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2635
diff
changeset
|
957 |
assumes a: "(p \<bullet> bs) \<sharp>* x" |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
958 |
(*and b: "finite bs"*) |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
959 |
shows "\<exists>q. [bs]set. x = [p \<bullet> bs]set. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
960 |
proof - |
3058 | 961 |
from set_renaming_perm2 |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
962 |
obtain q where *: "\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> bs \<union> (p \<bullet> bs)" by blast |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
963 |
have ***: "q \<bullet> bs = p \<bullet> bs" using * |
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
964 |
unfolding permute_set_eq_image image_def by auto |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
965 |
have "[bs]set. x = q \<bullet> ([bs]set. x)" |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
966 |
apply(rule perm_supp_eq[symmetric]) |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
967 |
using a ** |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
968 |
unfolding Abs_fresh_star_iff |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
969 |
unfolding fresh_star_def |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
970 |
by auto |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
971 |
also have "\<dots> = [q \<bullet> bs]set. (q \<bullet> x)" by simp |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
972 |
finally have "[bs]set. x = [p \<bullet> bs]set. (q \<bullet> x)" by (simp add: ***) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
973 |
then show "\<exists>q. [bs]set. x = [p \<bullet> bs]set. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" using *** by metis |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
974 |
qed |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
975 |
|
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
976 |
lemma Abs_rename_res: |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
977 |
fixes x::"'a::fs" |
2659
619ecb57db38
strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2635
diff
changeset
|
978 |
assumes a: "(p \<bullet> bs) \<sharp>* x" |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
979 |
(*and b: "finite bs"*) |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
980 |
shows "\<exists>q. [bs]res. x = [p \<bullet> bs]res. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
981 |
proof - |
3058 | 982 |
from set_renaming_perm2 |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
983 |
obtain q where *: "\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> bs \<union> (p \<bullet> bs)" by blast |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
984 |
have ***: "q \<bullet> bs = p \<bullet> bs" using * |
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
985 |
unfolding permute_set_eq_image image_def by auto |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
986 |
have "[bs]res. x = q \<bullet> ([bs]res. x)" |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
987 |
apply(rule perm_supp_eq[symmetric]) |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
988 |
using a ** |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
989 |
unfolding Abs_fresh_star_iff |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
990 |
unfolding fresh_star_def |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
991 |
by auto |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
992 |
also have "\<dots> = [q \<bullet> bs]res. (q \<bullet> x)" by simp |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
993 |
finally have "[bs]res. x = [p \<bullet> bs]res. (q \<bullet> x)" by (simp add: ***) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
994 |
then show "\<exists>q. [bs]res. x = [p \<bullet> bs]res. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" using *** by metis |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
995 |
qed |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
996 |
|
2611
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
Christian Urban <urbanc@in.tum.de>
parents:
2599
diff
changeset
|
997 |
lemma Abs_rename_lst: |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
998 |
fixes x::"'a::fs" |
2659
619ecb57db38
strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2635
diff
changeset
|
999 |
assumes a: "(p \<bullet> (set bs)) \<sharp>* x" |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1000 |
shows "\<exists>q. [bs]lst. x = [p \<bullet> bs]lst. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1001 |
proof - |
3058 | 1002 |
from list_renaming_perm |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
1003 |
obtain q where *: "\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> set bs \<union> (p \<bullet> set bs)" by blast |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
1004 |
have ***: "q \<bullet> bs = p \<bullet> bs" using * by (induct bs) (simp_all add: insert_eqvt) |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1005 |
have "[bs]lst. x = q \<bullet> ([bs]lst. x)" |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1006 |
apply(rule perm_supp_eq[symmetric]) |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1007 |
using a ** |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1008 |
unfolding Abs_fresh_star_iff |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1009 |
unfolding fresh_star_def |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1010 |
by auto |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1011 |
also have "\<dots> = [q \<bullet> bs]lst. (q \<bullet> x)" by simp |
2668
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
1012 |
finally have "[bs]lst. x = [p \<bullet> bs]lst. (q \<bullet> x)" by (simp add: ***) |
92c001d93225
modified the renaming_perm lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2663
diff
changeset
|
1013 |
then show "\<exists>q. [bs]lst. x = [p \<bullet> bs]lst. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" using *** by metis |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1014 |
qed |
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1015 |
|
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1016 |
|
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1017 |
text {* for deep recursive binders *} |
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1018 |
|
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1019 |
lemma Abs_rename_set': |
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1020 |
fixes x::"'a::fs" |
2659
619ecb57db38
strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2635
diff
changeset
|
1021 |
assumes a: "(p \<bullet> bs) \<sharp>* x" |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
1022 |
(*and b: "finite bs"*) |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1023 |
shows "\<exists>q. [bs]set. x = [q \<bullet> bs]set. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
1024 |
using Abs_rename_set[OF a] by metis |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1025 |
|
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1026 |
lemma Abs_rename_res': |
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1027 |
fixes x::"'a::fs" |
2659
619ecb57db38
strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2635
diff
changeset
|
1028 |
assumes a: "(p \<bullet> bs) \<sharp>* x" |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
1029 |
(*and b: "finite bs"*) |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1030 |
shows "\<exists>q. [bs]res. x = [q \<bullet> bs]res. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" |
3060
6613514ff6cb
tiny improvement by removing one unnecessary assumption
Christian Urban <urbanc@in.tum.de>
parents:
3058
diff
changeset
|
1031 |
using Abs_rename_res[OF a] by metis |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1032 |
|
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1033 |
lemma Abs_rename_lst': |
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1034 |
fixes x::"'a::fs" |
2659
619ecb57db38
strengthened renaming lemmas
Christian Urban <urbanc@in.tum.de>
parents:
2635
diff
changeset
|
1035 |
assumes a: "(p \<bullet> (set bs)) \<sharp>* x" |
2616
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1036 |
shows "\<exists>q. [bs]lst. x = [q \<bullet> bs]lst. (q \<bullet> x) \<and> q \<bullet> bs = p \<bullet> bs" |
dd7490fdd998
all examples for strong exhausts work; recursive binders need to be treated differently; still unclean version with lots of diagnostic code
Christian Urban <urbanc@in.tum.de>
parents:
2611
diff
changeset
|
1037 |
using Abs_rename_lst[OF a] by metis |
2599
d6fe94028a5d
moved general theorems into the libraries
Christian Urban <urbanc@in.tum.de>
parents:
2592
diff
changeset
|
1038 |
|
2468 | 1039 |
section {* Infrastructure for building tuples of relations and functions *} |
1040 |
||
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1041 |
fun |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1042 |
prod_fv :: "('a \<Rightarrow> atom set) \<Rightarrow> ('b \<Rightarrow> atom set) \<Rightarrow> ('a \<times> 'b) \<Rightarrow> atom set" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1043 |
where |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1044 |
"prod_fv fv1 fv2 (x, y) = fv1 x \<union> fv2 y" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1045 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1046 |
definition |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1047 |
prod_alpha :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool)" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1048 |
where |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1049 |
"prod_alpha = prod_rel" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1050 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1051 |
lemma [quot_respect]: |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1052 |
shows "((R1 ===> op =) ===> (R2 ===> op =) ===> prod_rel R1 R2 ===> op =) prod_fv prod_fv" |
2559
add799cf0817
adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents:
2491
diff
changeset
|
1053 |
unfolding fun_rel_def |
add799cf0817
adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents:
2491
diff
changeset
|
1054 |
unfolding prod_rel_def |
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1055 |
by auto |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1056 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1057 |
lemma [quot_preserve]: |
3157
de89c95c5377
updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
parents:
3152
diff
changeset
|
1058 |
assumes q1: "Quotient3 R1 abs1 rep1" |
de89c95c5377
updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
parents:
3152
diff
changeset
|
1059 |
and q2: "Quotient3 R2 abs2 rep2" |
2574 | 1060 |
shows "((abs1 ---> id) ---> (abs2 ---> id) ---> map_pair rep1 rep2 ---> id) prod_fv = prod_fv" |
3157
de89c95c5377
updated to latest changes (10 April) to quotient package (lift_raw_const only takes dummy theorem TrueI....in the future this will not work anymore)
Christian Urban <urbanc@in.tum.de>
parents:
3152
diff
changeset
|
1061 |
by (simp add: fun_eq_iff Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2]) |
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1062 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1063 |
lemma [mono]: |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1064 |
shows "A <= B \<Longrightarrow> C <= D ==> prod_alpha A C <= prod_alpha B D" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1065 |
unfolding prod_alpha_def |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1066 |
by auto |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1067 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1068 |
lemma [eqvt]: |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1069 |
shows "p \<bullet> prod_alpha A B x y = prod_alpha (p \<bullet> A) (p \<bullet> B) (p \<bullet> x) (p \<bullet> y)" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1070 |
unfolding prod_alpha_def |
2559
add799cf0817
adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents:
2491
diff
changeset
|
1071 |
unfolding prod_rel_def |
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1072 |
by (perm_simp) (rule refl) |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1073 |
|
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1074 |
lemma [eqvt]: |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1075 |
shows "p \<bullet> prod_fv A B (x, y) = prod_fv (p \<bullet> A) (p \<bullet> B) (p \<bullet> x, p \<bullet> y)" |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1076 |
unfolding prod_fv.simps |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1077 |
by (perm_simp) (rule refl) |
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1078 |
|
2489
c0695bb33fcd
added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents:
2479
diff
changeset
|
1079 |
lemma prod_fv_supp: |
c0695bb33fcd
added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents:
2479
diff
changeset
|
1080 |
shows "prod_fv supp supp = supp" |
c0695bb33fcd
added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents:
2479
diff
changeset
|
1081 |
by (rule ext) |
c0695bb33fcd
added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents:
2479
diff
changeset
|
1082 |
(auto simp add: prod_fv.simps supp_Pair) |
c0695bb33fcd
added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents:
2479
diff
changeset
|
1083 |
|
c0695bb33fcd
added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents:
2479
diff
changeset
|
1084 |
lemma prod_alpha_eq: |
c0695bb33fcd
added simp rules for prod_fv and prod_alpha
Christian Urban <urbanc@in.tum.de>
parents:
2479
diff
changeset
|
1085 |
shows "prod_alpha (op=) (op=) = (op=)" |
2843
1ae3c9b2d557
Slightly modify fcb for list1 and put in common place.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2733
diff
changeset
|
1086 |
unfolding prod_alpha_def |
1ae3c9b2d557
Slightly modify fcb for list1 and put in common place.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2733
diff
changeset
|
1087 |
by (auto intro!: ext) |
1ae3c9b2d557
Slightly modify fcb for list1 and put in common place.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2733
diff
changeset
|
1088 |
|
2385
fe25a3ffeb14
cleaned up a bit Abs.thy
Christian Urban <urbanc@in.tum.de>
parents:
2324
diff
changeset
|
1089 |
|
3199 | 1090 |
|
1091 |
||
1092 |
||
1093 |
||
1094 |
||
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1095 |
end |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1096 |