author | Christian Urban <urbanc@in.tum.de> |
Mon, 07 Jun 2010 15:13:39 +0200 | |
changeset 2214 | 02e03d4287ec |
parent 2124 | a17b25cb94a6 |
child 2296 | 45a69c9cc4cc |
permissions | -rw-r--r-- |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory Abs |
1804
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
2 |
imports "../Nominal-General/Nominal2_Atoms" |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
3 |
"../Nominal-General/Nominal2_Eqvt" |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
4 |
"../Nominal-General/Nominal2_Supp" |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
5 |
"Quotient" |
1933
9eab1dfc14d2
moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents:
1932
diff
changeset
|
6 |
"Quotient_List" |
1804
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
7 |
"Quotient_Product" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
begin |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
fun |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
alpha_gen |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
where |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
alpha_gen[simp del]: |
1465
4de35639fef0
added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1460
diff
changeset
|
14 |
"alpha_gen (bs, x) R f pi (cs, y) \<longleftrightarrow> |
4de35639fef0
added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1460
diff
changeset
|
15 |
f x - bs = f y - cs \<and> |
4de35639fef0
added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1460
diff
changeset
|
16 |
(f x - bs) \<sharp>* pi \<and> |
4de35639fef0
added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1460
diff
changeset
|
17 |
R (pi \<bullet> x) y \<and> |
4de35639fef0
added another supp-proof for the non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1460
diff
changeset
|
18 |
pi \<bullet> bs = cs" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
|
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
20 |
fun |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
21 |
alpha_res |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
22 |
where |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
23 |
alpha_res[simp del]: |
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 (bs, x) R f pi (cs, y) \<longleftrightarrow> |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
25 |
f x - bs = f y - cs \<and> |
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) \<sharp>* pi \<and> |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
27 |
R (pi \<bullet> x) y" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
|
1557
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
29 |
fun |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
30 |
alpha_lst |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
31 |
where |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
32 |
alpha_lst[simp del]: |
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 (bs, x) R f pi (cs, y) \<longleftrightarrow> |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
34 |
f x - set bs = f y - set cs \<and> |
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) \<sharp>* pi \<and> |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
36 |
R (pi \<bullet> x) y \<and> |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
37 |
pi \<bullet> bs = cs" |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
38 |
|
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
39 |
lemmas alphas = alpha_gen.simps alpha_res.simps alpha_lst.simps |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
40 |
|
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
41 |
notation |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
42 |
alpha_gen ("_ \<approx>gen _ _ _ _" [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
|
43 |
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
|
44 |
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
|
45 |
|
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
46 |
(* monos *) |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
47 |
lemma [mono]: |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
48 |
shows "R1 \<le> R2 \<Longrightarrow> alpha_gen bs R1 \<le> alpha_gen bs R2" |
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
Christian Urban <urbanc@in.tum.de>
parents:
1544
diff
changeset
|
49 |
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
|
50 |
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
|
51 |
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
|
52 |
(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
|
53 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
fun |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
alpha_abs |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
where |
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
57 |
[simp del]: |
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
58 |
"alpha_abs (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op=) supp p (cs, y))" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
|
1657 | 60 |
fun |
61 |
alpha_abs_lst |
|
62 |
where |
|
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
63 |
[simp del]: |
1657 | 64 |
"alpha_abs_lst (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>lst (op=) supp p (cs, y))" |
65 |
||
66 |
fun |
|
67 |
alpha_abs_res |
|
68 |
where |
|
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
69 |
[simp del]: |
1657 | 70 |
"alpha_abs_res (bs, x) (cs, y) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op=) supp p (cs, y))" |
71 |
||
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
notation |
1666
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
73 |
alpha_abs (infix "\<approx>abs" 50) and |
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
74 |
alpha_abs_lst (infix "\<approx>abs'_lst" 50) and |
a99ae705b811
Removed some warnings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1665
diff
changeset
|
75 |
alpha_abs_res (infix "\<approx>abs'_res" 50) |
1657 | 76 |
|
77 |
lemmas alphas_abs = alpha_abs.simps alpha_abs_res.simps alpha_abs_lst.simps |
|
78 |
||
79 |
lemma alphas_abs_refl: |
|
80 |
shows "(bs, x) \<approx>abs (bs, x)" |
|
81 |
and "(bs, x) \<approx>abs_res (bs, x)" |
|
82 |
and "(cs, x) \<approx>abs_lst (cs, x)" |
|
83 |
unfolding alphas_abs |
|
84 |
unfolding alphas |
|
85 |
unfolding fresh_star_def |
|
86 |
by (rule_tac [!] x="0" in exI) |
|
87 |
(simp_all add: fresh_zero_perm) |
|
88 |
||
89 |
lemma alphas_abs_sym: |
|
90 |
shows "(bs, x) \<approx>abs (cs, y) \<Longrightarrow> (cs, y) \<approx>abs (bs, x)" |
|
91 |
and "(bs, x) \<approx>abs_res (cs, y) \<Longrightarrow> (cs, y) \<approx>abs_res (bs, x)" |
|
92 |
and "(ds, x) \<approx>abs_lst (es, y) \<Longrightarrow> (es, y) \<approx>abs_lst (ds, x)" |
|
93 |
unfolding alphas_abs |
|
94 |
unfolding alphas |
|
95 |
unfolding fresh_star_def |
|
96 |
by (erule_tac [!] exE, rule_tac [!] x="-p" in exI) |
|
97 |
(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
|
98 |
|
1657 | 99 |
lemma alphas_abs_trans: |
100 |
shows "\<lbrakk>(bs, x) \<approx>abs (cs, y); (cs, y) \<approx>abs (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>abs (ds, z)" |
|
101 |
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)" |
|
102 |
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)" |
|
103 |
unfolding alphas_abs |
|
104 |
unfolding alphas |
|
105 |
unfolding fresh_star_def |
|
106 |
apply(erule_tac [!] exE, erule_tac [!] exE) |
|
107 |
apply(rule_tac [!] x="pa + p" in exI) |
|
108 |
by (simp_all add: fresh_plus_perm) |
|
109 |
||
110 |
lemma alphas_abs_eqvt: |
|
111 |
shows "(bs, x) \<approx>abs (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>abs (p \<bullet> cs, p \<bullet> y)" |
|
112 |
and "(bs, x) \<approx>abs_res (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>abs_res (p \<bullet> cs, p \<bullet> y)" |
|
113 |
and "(ds, x) \<approx>abs_lst (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>abs_lst (p \<bullet> es, p \<bullet> y)" |
|
114 |
unfolding alphas_abs |
|
115 |
unfolding alphas |
|
116 |
unfolding set_eqvt[symmetric] |
|
117 |
unfolding supp_eqvt[symmetric] |
|
118 |
unfolding Diff_eqvt[symmetric] |
|
119 |
apply(erule_tac [!] exE) |
|
120 |
apply(rule_tac [!] x="p \<bullet> pa" in exI) |
|
121 |
by (auto simp add: fresh_star_permute_iff permute_eqvt[symmetric]) |
|
122 |
||
123 |
quotient_type |
|
124 |
'a abs_gen = "(atom set \<times> 'a::pt)" / "alpha_abs" |
|
125 |
and 'b abs_res = "(atom set \<times> 'b::pt)" / "alpha_abs_res" |
|
126 |
and 'c abs_lst = "(atom list \<times> 'c::pt)" / "alpha_abs_lst" |
|
127 |
apply(rule_tac [!] equivpI) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
unfolding reflp_def symp_def transp_def |
1657 | 129 |
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
|
130 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
131 |
quotient_definition |
1932 | 132 |
Abs ("[_]set. _" [60, 60] 60) |
133 |
where |
|
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
134 |
"Abs::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_gen" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
is |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
"Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)" |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
|
1657 | 138 |
quotient_definition |
1932 | 139 |
Abs_res ("[_]res. _" [60, 60] 60) |
140 |
where |
|
1657 | 141 |
"Abs_res::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_res" |
142 |
is |
|
143 |
"Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)" |
|
144 |
||
145 |
quotient_definition |
|
1932 | 146 |
Abs_lst ("[_]lst. _" [60, 60] 60) |
147 |
where |
|
1657 | 148 |
"Abs_lst::atom list \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs_lst" |
149 |
is |
|
150 |
"Pair::atom list \<Rightarrow> ('a::pt) \<Rightarrow> (atom list \<times> 'a)" |
|
151 |
||
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
152 |
lemma [quot_respect]: |
1657 | 153 |
shows "(op= ===> op= ===> alpha_abs) Pair Pair" |
154 |
and "(op= ===> op= ===> alpha_abs_res) Pair Pair" |
|
155 |
and "(op= ===> op= ===> alpha_abs_lst) Pair Pair" |
|
156 |
unfolding fun_rel_def |
|
157 |
by (auto intro: alphas_abs_refl simp only:) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
158 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
159 |
lemma [quot_respect]: |
1657 | 160 |
shows "(op= ===> alpha_abs ===> alpha_abs) permute permute" |
161 |
and "(op= ===> alpha_abs_res ===> alpha_abs_res) permute permute" |
|
162 |
and "(op= ===> alpha_abs_lst ===> alpha_abs_lst) permute permute" |
|
163 |
unfolding fun_rel_def |
|
164 |
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
|
165 |
|
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
|
166 |
lemma abs_exhausts: |
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
|
167 |
shows "(\<And>as (x::'a::pt). y1 = Abs as x \<Longrightarrow> P1) \<Longrightarrow> P1" |
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
|
168 |
and "(\<And>as (x::'a::pt). y2 = Abs_res as x \<Longrightarrow> P2) \<Longrightarrow> P2" |
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
|
169 |
and "(\<And>as (x::'a::pt). y3 = Abs_lst as x \<Longrightarrow> P3) \<Longrightarrow> P3" |
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
|
170 |
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
|
171 |
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
|
172 |
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
|
173 |
|
1662 | 174 |
lemma abs_eq_iff: |
175 |
shows "Abs bs x = Abs cs y \<longleftrightarrow> (bs, x) \<approx>abs (cs, y)" |
|
176 |
and "Abs_res bs x = Abs_res cs y \<longleftrightarrow> (bs, x) \<approx>abs_res (cs, y)" |
|
177 |
and "Abs_lst ds x = Abs_lst hs y \<longleftrightarrow> (ds, x) \<approx>abs_lst (hs, y)" |
|
1932 | 178 |
unfolding alphas_abs |
179 |
by (lifting alphas_abs) |
|
1662 | 180 |
|
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
181 |
instantiation abs_gen :: (pt) pt |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
begin |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
quotient_definition |
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
185 |
"permute_abs_gen::perm \<Rightarrow> ('a::pt abs_gen) \<Rightarrow> 'a abs_gen" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
186 |
is |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
187 |
"permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)" |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
188 |
|
1657 | 189 |
lemma permute_Abs[simp]: |
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
190 |
fixes x::"'a::pt" |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
191 |
shows "(p \<bullet> (Abs as x)) = Abs (p \<bullet> as) (p \<bullet> x)" |
1657 | 192 |
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
|
193 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
194 |
instance |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
195 |
apply(default) |
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
|
196 |
apply(case_tac [!] x rule: abs_exhausts(1)) |
1657 | 197 |
apply(simp_all) |
198 |
done |
|
199 |
||
200 |
end |
|
201 |
||
202 |
instantiation abs_res :: (pt) pt |
|
203 |
begin |
|
204 |
||
205 |
quotient_definition |
|
206 |
"permute_abs_res::perm \<Rightarrow> ('a::pt abs_res) \<Rightarrow> 'a abs_res" |
|
207 |
is |
|
208 |
"permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)" |
|
209 |
||
210 |
lemma permute_Abs_res[simp]: |
|
211 |
fixes x::"'a::pt" |
|
212 |
shows "(p \<bullet> (Abs_res as x)) = Abs_res (p \<bullet> as) (p \<bullet> x)" |
|
213 |
by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"]) |
|
214 |
||
215 |
instance |
|
216 |
apply(default) |
|
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
|
217 |
apply(case_tac [!] x rule: abs_exhausts(2)) |
1657 | 218 |
apply(simp_all) |
219 |
done |
|
220 |
||
221 |
end |
|
222 |
||
223 |
instantiation abs_lst :: (pt) pt |
|
224 |
begin |
|
225 |
||
226 |
quotient_definition |
|
227 |
"permute_abs_lst::perm \<Rightarrow> ('a::pt abs_lst) \<Rightarrow> 'a abs_lst" |
|
228 |
is |
|
229 |
"permute:: perm \<Rightarrow> (atom list \<times> 'a::pt) \<Rightarrow> (atom list \<times> 'a::pt)" |
|
230 |
||
231 |
lemma permute_Abs_lst[simp]: |
|
232 |
fixes x::"'a::pt" |
|
233 |
shows "(p \<bullet> (Abs_lst as x)) = Abs_lst (p \<bullet> as) (p \<bullet> x)" |
|
234 |
by (lifting permute_prod.simps[where 'a="atom list" and 'b="'a"]) |
|
235 |
||
236 |
instance |
|
237 |
apply(default) |
|
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
|
238 |
apply(case_tac [!] x rule: abs_exhausts(3)) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
239 |
apply(simp_all) |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
240 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
241 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
242 |
end |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
243 |
|
1657 | 244 |
lemmas permute_abs = permute_Abs permute_Abs_res permute_Abs_lst |
245 |
||
1662 | 246 |
lemma abs_swap1: |
247 |
assumes a1: "a \<notin> (supp x) - bs" |
|
248 |
and a2: "b \<notin> (supp x) - bs" |
|
249 |
shows "Abs bs x = Abs ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x)" |
|
250 |
and "Abs_res bs x = Abs_res ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x)" |
|
251 |
unfolding abs_eq_iff |
|
252 |
unfolding alphas_abs |
|
253 |
unfolding alphas |
|
254 |
unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] |
|
255 |
unfolding fresh_star_def fresh_def |
|
256 |
unfolding swap_set_not_in[OF a1 a2] |
|
257 |
using a1 a2 |
|
258 |
by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI) |
|
259 |
(auto simp add: supp_perm swap_atom) |
|
260 |
||
261 |
lemma abs_swap2: |
|
262 |
assumes a1: "a \<notin> (supp x) - (set bs)" |
|
263 |
and a2: "b \<notin> (supp x) - (set bs)" |
|
264 |
shows "Abs_lst bs x = Abs_lst ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x)" |
|
265 |
unfolding abs_eq_iff |
|
266 |
unfolding alphas_abs |
|
267 |
unfolding alphas |
|
268 |
unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] set_eqvt[symmetric] |
|
269 |
unfolding fresh_star_def fresh_def |
|
270 |
unfolding swap_set_not_in[OF a1 a2] |
|
271 |
using a1 a2 |
|
272 |
by (rule_tac [!] x="(a \<rightleftharpoons> b)" in exI) |
|
273 |
(auto simp add: supp_perm swap_atom) |
|
274 |
||
275 |
lemma abs_supports: |
|
276 |
shows "((supp x) - as) supports (Abs as x)" |
|
277 |
and "((supp x) - as) supports (Abs_res as x)" |
|
278 |
and "((supp x) - (set bs)) supports (Abs_lst bs x)" |
|
279 |
unfolding supports_def |
|
280 |
unfolding permute_abs |
|
281 |
by (simp_all add: abs_swap1[symmetric] abs_swap2[symmetric]) |
|
1657 | 282 |
|
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
|
283 |
function |
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
|
284 |
supp_gen :: "('a::pt) abs_gen \<Rightarrow> atom set" |
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
|
285 |
where |
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
|
286 |
"supp_gen (Abs as x) = supp x - as" |
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
|
287 |
apply(case_tac x rule: abs_exhausts(1)) |
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
|
288 |
apply(simp) |
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
|
289 |
apply(simp add: abs_eq_iff alphas_abs alphas) |
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
|
290 |
done |
1657 | 291 |
|
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
|
292 |
termination supp_gen |
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
|
293 |
by (auto intro!: local.termination) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
294 |
|
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
|
295 |
function |
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
|
296 |
supp_res :: "('a::pt) abs_res \<Rightarrow> atom set" |
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
|
297 |
where |
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
|
298 |
"supp_res (Abs_res as x) = supp x - as" |
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
|
299 |
apply(case_tac x rule: abs_exhausts(2)) |
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
|
300 |
apply(simp) |
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
|
301 |
apply(simp add: abs_eq_iff alphas_abs alphas) |
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
|
302 |
done |
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
|
303 |
|
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
|
304 |
termination supp_res |
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
|
305 |
by (auto intro!: local.termination) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
306 |
|
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
|
307 |
function |
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
|
308 |
supp_lst :: "('a::pt) abs_lst \<Rightarrow> atom set" |
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
|
309 |
where |
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
|
310 |
"supp_lst (Abs_lst cs x) = (supp x) - (set cs)" |
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
|
311 |
apply(case_tac x rule: abs_exhausts(3)) |
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
|
312 |
apply(simp) |
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
|
313 |
apply(simp add: abs_eq_iff alphas_abs alphas) |
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
|
314 |
done |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
315 |
|
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
|
316 |
termination supp_lst |
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
|
317 |
by (auto intro!: local.termination) |
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
|
318 |
|
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
|
319 |
lemma [eqvt]: |
1657 | 320 |
shows "(p \<bullet> supp_gen x) = supp_gen (p \<bullet> x)" |
321 |
and "(p \<bullet> supp_res y) = supp_res (p \<bullet> y)" |
|
322 |
and "(p \<bullet> supp_lst z) = supp_lst (p \<bullet> z)" |
|
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
|
323 |
apply(case_tac x rule: abs_exhausts(1)) |
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
|
324 |
apply(simp add: supp_eqvt Diff_eqvt) |
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
|
325 |
apply(case_tac y rule: abs_exhausts(2)) |
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
|
326 |
apply(simp add: supp_eqvt Diff_eqvt) |
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
|
327 |
apply(case_tac z rule: abs_exhausts(3)) |
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
|
328 |
apply(simp add: supp_eqvt Diff_eqvt set_eqvt) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
329 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
330 |
|
1657 | 331 |
lemma aux_fresh: |
332 |
shows "a \<sharp> Abs bs x \<Longrightarrow> a \<sharp> supp_gen (Abs bs x)" |
|
333 |
and "a \<sharp> Abs_res bs x \<Longrightarrow> a \<sharp> supp_res (Abs_res bs x)" |
|
334 |
and "a \<sharp> Abs_lst cs x \<Longrightarrow> a \<sharp> supp_lst (Abs_lst cs x)" |
|
1932 | 335 |
by (rule_tac [!] fresh_fun_eqvt_app) |
336 |
(simp_all add: eqvts_raw) |
|
1657 | 337 |
|
338 |
lemma supp_abs_subset1: |
|
339 |
assumes a: "finite (supp x)" |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
340 |
shows "(supp x) - as \<subseteq> supp (Abs as x)" |
1657 | 341 |
and "(supp x) - as \<subseteq> supp (Abs_res as x)" |
342 |
and "(supp x) - (set bs) \<subseteq> supp (Abs_lst bs x)" |
|
343 |
unfolding supp_conv_fresh |
|
1932 | 344 |
by (auto dest!: aux_fresh) |
345 |
(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
|
346 |
|
1657 | 347 |
lemma supp_abs_subset2: |
348 |
assumes a: "finite (supp x)" |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
349 |
shows "supp (Abs as x) \<subseteq> (supp x) - as" |
1657 | 350 |
and "supp (Abs_res as x) \<subseteq> (supp x) - as" |
351 |
and "supp (Abs_lst bs x) \<subseteq> (supp x) - (set bs)" |
|
1932 | 352 |
by (rule_tac [!] supp_is_subset) |
353 |
(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
|
354 |
|
1657 | 355 |
lemma abs_finite_supp: |
356 |
assumes a: "finite (supp x)" |
|
1478
1ea4ca823266
added proof of supp/fv for type schemes
Christian Urban <urbanc@in.tum.de>
parents:
1470
diff
changeset
|
357 |
shows "supp (Abs as x) = (supp x) - as" |
1657 | 358 |
and "supp (Abs_res as x) = (supp x) - as" |
359 |
and "supp (Abs_lst bs x) = (supp x) - (set bs)" |
|
1932 | 360 |
by (rule_tac [!] subset_antisym) |
361 |
(simp_all add: supp_abs_subset1[OF a] supp_abs_subset2[OF a]) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
362 |
|
1657 | 363 |
lemma supp_abs: |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
364 |
fixes x::"'a::fs" |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
365 |
shows "supp (Abs as x) = (supp x) - as" |
1657 | 366 |
and "supp (Abs_res as x) = (supp x) - as" |
367 |
and "supp (Abs_lst bs x) = (supp x) - (set bs)" |
|
1932 | 368 |
by (rule_tac [!] abs_finite_supp) |
369 |
(simp_all add: finite_supp) |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
370 |
|
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
371 |
instance abs_gen :: (fs) fs |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
372 |
apply(default) |
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
|
373 |
apply(case_tac x rule: abs_exhausts(1)) |
1657 | 374 |
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
|
375 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
376 |
|
1657 | 377 |
instance abs_res :: (fs) fs |
378 |
apply(default) |
|
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
|
379 |
apply(case_tac x rule: abs_exhausts(2)) |
1657 | 380 |
apply(simp add: supp_abs finite_supp) |
381 |
done |
|
382 |
||
383 |
instance abs_lst :: (fs) fs |
|
384 |
apply(default) |
|
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
|
385 |
apply(case_tac x rule: abs_exhausts(3)) |
1657 | 386 |
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
|
387 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
388 |
|
1657 | 389 |
lemma abs_fresh_iff: |
390 |
fixes x::"'a::fs" |
|
391 |
shows "a \<sharp> Abs bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)" |
|
392 |
and "a \<sharp> Abs_res bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)" |
|
393 |
and "a \<sharp> Abs_lst cs x \<longleftrightarrow> a \<in> (set cs) \<or> (a \<notin> (set cs) \<and> a \<sharp> x)" |
|
394 |
unfolding fresh_def |
|
395 |
unfolding supp_abs |
|
396 |
by auto |
|
1460
0fd03936dedb
merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1451
diff
changeset
|
397 |
|
1657 | 398 |
section {* BELOW is stuff that may or may not be needed *} |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
399 |
|
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
400 |
lemma supp_atom_image: |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
401 |
fixes as::"'a::at_base set" |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
402 |
shows "supp (atom ` as) = supp as" |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
403 |
apply(simp add: supp_def) |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
404 |
apply(simp add: image_eqvt) |
1933
9eab1dfc14d2
moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents:
1932
diff
changeset
|
405 |
apply(simp add: eqvts_raw) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
406 |
apply(simp add: atom_image_cong) |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
407 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
408 |
|
1460
0fd03936dedb
merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1451
diff
changeset
|
409 |
lemma swap_atom_image_fresh: "\<lbrakk>a \<sharp> atom ` (fn :: ('a :: at_base set)); b \<sharp> atom ` fn\<rbrakk> \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> fn = fn" |
0fd03936dedb
merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1451
diff
changeset
|
410 |
apply (simp add: fresh_def) |
0fd03936dedb
merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1451
diff
changeset
|
411 |
apply (simp add: supp_atom_image) |
0fd03936dedb
merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1451
diff
changeset
|
412 |
apply (fold fresh_def) |
0fd03936dedb
merge and proof of support for non-recursive case
Christian Urban <urbanc@in.tum.de>
parents:
1451
diff
changeset
|
413 |
apply (simp add: swap_fresh_fresh) |
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
414 |
done |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
415 |
|
1467
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
416 |
(* TODO: The following lemmas can be moved somewhere... *) |
1657 | 417 |
|
418 |
lemma Abs_eq_iff: |
|
419 |
shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))" |
|
1675
d24f59f78a86
Generalize Abs_eq_iff.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
420 |
and "Abs_res bs x = Abs_res cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op =) supp p (cs, y))" |
d24f59f78a86
Generalize Abs_eq_iff.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
421 |
and "Abs_lst bsl x = Abs_lst csl y \<longleftrightarrow> (\<exists>p. (bsl, x) \<approx>lst (op =) supp p (csl, y))" |
d24f59f78a86
Generalize Abs_eq_iff.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
422 |
by (lifting alphas_abs) |
1657 | 423 |
|
1467
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
424 |
lemma split_rsp2[quot_respect]: "((R1 ===> R2 ===> prod_rel R1 R2 ===> op =) ===> |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
425 |
prod_rel R1 R2 ===> prod_rel R1 R2 ===> op =) split split" |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
426 |
by auto |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
427 |
|
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
428 |
lemma split_prs2[quot_preserve]: |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
429 |
assumes q1: "Quotient R1 Abs1 Rep1" |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
430 |
and q2: "Quotient R2 Abs2 Rep2" |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
431 |
shows "((Abs1 ---> Abs2 ---> prod_fun Abs1 Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> prod_fun Rep1 Rep2 ---> id) split = split" |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
432 |
by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) |
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
433 |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
434 |
lemma alphas2: |
1467
77b86f1fc936
Lifting theorems with compound fv and compound alpha.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1460
diff
changeset
|
435 |
"(bs, x1, x2) \<approx>gen (\<lambda>(x1, y1) (x2, y2). R1 x1 x2 \<and> R2 y1 y2) (\<lambda>(a, b). f1 a \<union> f2 b) pi (cs, y1, y2) = |
1470
3127c75275a6
Fix for the change of alpha_gen.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1469
diff
changeset
|
436 |
(f1 x1 \<union> f2 x2 - bs = f1 y1 \<union> f2 y2 - cs \<and> (f1 x1 \<union> f2 x2 - bs) \<sharp>* pi \<and> R1 (pi \<bullet> x1) y1 \<and> R2 (pi \<bullet> x2) y2 |
3127c75275a6
Fix for the change of alpha_gen.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1469
diff
changeset
|
437 |
\<and> pi \<bullet> bs = cs)" |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
438 |
"(bs, x1, x2) \<approx>res (\<lambda>(x1, y1) (x2, y2). R1 x1 x2 \<and> R2 y1 y2) (\<lambda>(a, b). f1 a \<union> f2 b) pi (cs, y1, y2) = |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
439 |
(f1 x1 \<union> f2 x2 - bs = f1 y1 \<union> f2 y2 - cs \<and> (f1 x1 \<union> f2 x2 - bs) \<sharp>* pi \<and> R1 (pi \<bullet> x1) y1 \<and> R2 (pi \<bullet> x2) y2)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
440 |
"(bsl, x1, x2) \<approx>lst (\<lambda>(x1, y1) (x2, y2). R1 x1 x2 \<and> R2 y1 y2) (\<lambda>(a, b). f1 a \<union> f2 b) pi (csl, y1, y2) = |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
441 |
(f1 x1 \<union> f2 x2 - set bsl = f1 y1 \<union> f2 y2 - set csl \<and> (f1 x1 \<union> f2 x2 - set bsl) \<sharp>* pi \<and> R1 (pi \<bullet> x1) y1 \<and> R2 (pi \<bullet> x2) y2 |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
442 |
\<and> pi \<bullet> bsl = csl)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
443 |
by (simp_all add: alphas) |
1544
c6849a634582
Keep only one copy of infinite_Un.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1543
diff
changeset
|
444 |
|
1744
00680cea0dde
Let with multiple bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1691
diff
changeset
|
445 |
lemma alphas3: |
00680cea0dde
Let with multiple bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1691
diff
changeset
|
446 |
"(bsl, x1, x2, x3) \<approx>lst (\<lambda>(x1, y1, z1) (x2, y2, z2). R1 x1 x2 \<and> R2 y1 y2 \<and> R3 z1 z2) (\<lambda>(a, b, c). f1 a \<union> (f2 b \<union> f3 c)) pi (csl, y1, y2, y3) = (f1 x1 \<union> (f2 x2 \<union> f3 x3) - set bsl = f1 y1 \<union> (f2 y2 \<union> f3 y3) - set csl \<and> |
00680cea0dde
Let with multiple bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1691
diff
changeset
|
447 |
(f1 x1 \<union> (f2 x2 \<union> f3 x3) - set bsl) \<sharp>* pi \<and> |
00680cea0dde
Let with multiple bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1691
diff
changeset
|
448 |
R1 (pi \<bullet> x1) y1 \<and> R2 (pi \<bullet> x2) y2 \<and> R3 (pi \<bullet> x3) y3 \<and> pi \<bullet> bsl = csl)" |
00680cea0dde
Let with multiple bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1691
diff
changeset
|
449 |
by (simp add: alphas) |
00680cea0dde
Let with multiple bindings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1691
diff
changeset
|
450 |
|
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
451 |
lemma alpha_gen_compose_sym: |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
452 |
fixes pi |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
453 |
assumes b: "(aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi (ab, s)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
454 |
and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
455 |
shows "(ab, s) \<approx>gen R f (- pi) (aa, t)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
456 |
using b apply - |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
457 |
apply(simp add: alphas) |
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
458 |
apply(erule conjE)+ |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
459 |
apply(rule conjI) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
460 |
apply(simp add: fresh_star_def fresh_minus_perm) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
461 |
apply(subgoal_tac "R (- pi \<bullet> s) ((- pi) \<bullet> (pi \<bullet> t))") |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
462 |
apply simp |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
463 |
apply(clarify) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
464 |
apply(simp) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
465 |
apply(rule a) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
466 |
apply assumption |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
467 |
done |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
468 |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
469 |
lemma alpha_res_compose_sym: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
470 |
fixes pi |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
471 |
assumes b: "(aa, t) \<approx>res (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi (ab, s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
472 |
and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
473 |
shows "(ab, s) \<approx>res R f (- pi) (aa, t)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
474 |
using b apply - |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
475 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
476 |
apply(erule conjE)+ |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
477 |
apply(rule conjI) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
478 |
apply(simp add: fresh_star_def fresh_minus_perm) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
479 |
apply(subgoal_tac "R (- pi \<bullet> s) ((- pi) \<bullet> (pi \<bullet> t))") |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
480 |
apply simp |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
481 |
apply(rule a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
482 |
apply assumption |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
483 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
484 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
485 |
lemma alpha_lst_compose_sym: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
486 |
fixes pi |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
487 |
assumes b: "(aa, t) \<approx>lst (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi (ab, s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
488 |
and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
489 |
shows "(ab, s) \<approx>lst R f (- pi) (aa, t)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
490 |
using b apply - |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
491 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
492 |
apply(erule conjE)+ |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
493 |
apply(rule conjI) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
494 |
apply(simp add: fresh_star_def fresh_minus_perm) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
495 |
apply(subgoal_tac "R (- pi \<bullet> s) ((- pi) \<bullet> (pi \<bullet> t))") |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
496 |
apply simp |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
497 |
apply(clarify) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
498 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
499 |
apply(rule a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
500 |
apply assumption |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
501 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
502 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
503 |
lemmas alphas_compose_sym = alpha_gen_compose_sym alpha_res_compose_sym alpha_lst_compose_sym |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
504 |
|
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
505 |
lemma alpha_gen_compose_sym2: |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
506 |
assumes a: "(aa, t1, t2) \<approx>gen (\<lambda>(x11, x12) (x21, x22). |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
507 |
(R1 x11 x21 \<and> R1 x21 x11) \<and> R2 x12 x22 \<and> R2 x22 x12) (\<lambda>(b, a). fb b \<union> fa a) pi (ab, s1, s2)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
508 |
and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
509 |
and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
510 |
shows "(ab, s1, s2) \<approx>gen (\<lambda>(a, b) (d, c). R1 a d \<and> R2 b c) (\<lambda>(b, a). fb b \<union> fa a) (- pi) (aa, t1, t2)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
511 |
using a |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
512 |
apply(simp add: alphas) |
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
513 |
apply clarify |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
514 |
apply (rule conjI) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
515 |
apply(simp add: fresh_star_def fresh_minus_perm) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
516 |
apply (rule conjI) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
517 |
apply (rotate_tac 3) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
518 |
apply (drule_tac pi="- pi" in r1) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
519 |
apply simp |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
520 |
apply (rule conjI) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
521 |
apply (rotate_tac -1) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
522 |
apply (drule_tac pi="- pi" in r2) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
523 |
apply simp_all |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
524 |
done |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
525 |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
526 |
lemma alpha_res_compose_sym2: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
527 |
assumes a: "(aa, t1, t2) \<approx>res (\<lambda>(x11, x12) (x21, x22). |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
528 |
(R1 x11 x21 \<and> R1 x21 x11) \<and> R2 x12 x22 \<and> R2 x22 x12) (\<lambda>(b, a). fb b \<union> fa a) pi (ab, s1, s2)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
529 |
and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
530 |
and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
531 |
shows "(ab, s1, s2) \<approx>res (\<lambda>(a, b) (d, c). R1 a d \<and> R2 b c) (\<lambda>(b, a). fb b \<union> fa a) (- pi) (aa, t1, t2)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
532 |
using a |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
533 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
534 |
apply clarify |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
535 |
apply (rule conjI) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
536 |
apply(simp add: fresh_star_def fresh_minus_perm) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
537 |
apply (rule conjI) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
538 |
apply (rotate_tac 3) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
539 |
apply (drule_tac pi="- pi" in r1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
540 |
apply simp |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
541 |
apply (rotate_tac -1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
542 |
apply (drule_tac pi="- pi" in r2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
543 |
apply simp |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
544 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
545 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
546 |
lemma alpha_lst_compose_sym2: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
547 |
assumes a: "(aa, t1, t2) \<approx>lst (\<lambda>(x11, x12) (x21, x22). |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
548 |
(R1 x11 x21 \<and> R1 x21 x11) \<and> R2 x12 x22 \<and> R2 x22 x12) (\<lambda>(b, a). fb b \<union> fa a) pi (ab, s1, s2)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
549 |
and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
550 |
and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
551 |
shows "(ab, s1, s2) \<approx>lst (\<lambda>(a, b) (d, c). R1 a d \<and> R2 b c) (\<lambda>(b, a). fb b \<union> fa a) (- pi) (aa, t1, t2)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
552 |
using a |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
553 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
554 |
apply clarify |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
555 |
apply (rule conjI) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
556 |
apply(simp add: fresh_star_def fresh_minus_perm) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
557 |
apply (rule conjI) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
558 |
apply (rotate_tac 3) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
559 |
apply (drule_tac pi="- pi" in r1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
560 |
apply simp |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
561 |
apply (rule conjI) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
562 |
apply (rotate_tac -1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
563 |
apply (drule_tac pi="- pi" in r2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
564 |
apply simp_all |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
565 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
566 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
567 |
lemmas alphas_compose_sym2 = alpha_gen_compose_sym2 alpha_res_compose_sym2 alpha_lst_compose_sym2 |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
568 |
|
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
569 |
lemma alpha_gen_compose_trans: |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
570 |
fixes pi pia |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
571 |
assumes b: "(aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi (ab, ta)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
572 |
and c: "(ab, ta) \<approx>gen R f pia (ac, sa)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
573 |
and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
574 |
shows "(aa, t) \<approx>gen R f (pia + pi) (ac, sa)" |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
575 |
using b c apply - |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
576 |
apply(simp add: alphas) |
1558
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
577 |
apply(erule conjE)+ |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
578 |
apply(simp add: fresh_star_plus) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
579 |
apply(drule_tac x="- pia \<bullet> sa" in spec) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
580 |
apply(drule mp) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
581 |
apply(rotate_tac 5) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
582 |
apply(drule_tac pi="- pia" in a) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
583 |
apply(simp) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
584 |
apply(rotate_tac 7) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
585 |
apply(drule_tac pi="pia" in a) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
586 |
apply(simp) |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
587 |
done |
a5ba76208983
started cleaning up and introduced 3 versions of ~~gen
Christian Urban <urbanc@in.tum.de>
parents:
1557
diff
changeset
|
588 |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
589 |
lemma alpha_res_compose_trans: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
590 |
fixes pi pia |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
591 |
assumes b: "(aa, t) \<approx>res (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi (ab, ta)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
592 |
and c: "(ab, ta) \<approx>res R f pia (ac, sa)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
593 |
and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
594 |
shows "(aa, t) \<approx>res R f (pia + pi) (ac, sa)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
595 |
using b c apply - |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
596 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
597 |
apply(erule conjE)+ |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
598 |
apply(simp add: fresh_star_plus) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
599 |
apply(drule_tac x="- pia \<bullet> sa" in spec) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
600 |
apply(drule mp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
601 |
apply(drule_tac pi="- pia" in a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
602 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
603 |
apply(rotate_tac 6) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
604 |
apply(drule_tac pi="pia" in a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
605 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
606 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
607 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
608 |
lemma alpha_lst_compose_trans: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
609 |
fixes pi pia |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
610 |
assumes b: "(aa, t) \<approx>lst (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi (ab, ta)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
611 |
and c: "(ab, ta) \<approx>lst R f pia (ac, sa)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
612 |
and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
613 |
shows "(aa, t) \<approx>lst R f (pia + pi) (ac, sa)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
614 |
using b c apply - |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
615 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
616 |
apply(erule conjE)+ |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
617 |
apply(simp add: fresh_star_plus) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
618 |
apply(drule_tac x="- pia \<bullet> sa" in spec) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
619 |
apply(drule mp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
620 |
apply(rotate_tac 5) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
621 |
apply(drule_tac pi="- pia" in a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
622 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
623 |
apply(rotate_tac 7) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
624 |
apply(drule_tac pi="pia" in a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
625 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
626 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
627 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
628 |
lemmas alphas_compose_trans = alpha_gen_compose_trans alpha_res_compose_trans alpha_lst_compose_trans |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
629 |
|
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
630 |
lemma alpha_gen_compose_trans2: |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
631 |
fixes pi pia |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
632 |
assumes b: "(aa, (t1, t2)) \<approx>gen |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
633 |
(\<lambda>(b, a) (d, c). R1 b d \<and> (\<forall>z. R1 d z \<longrightarrow> R1 b z) \<and> R2 a c \<and> (\<forall>z. R2 c z \<longrightarrow> R2 a z)) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
634 |
(\<lambda>(b, a). fv_a b \<union> fv_b a) pi (ab, (ta1, ta2))" |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
635 |
and c: "(ab, (ta1, ta2)) \<approx>gen (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
636 |
pia (ac, (sa1, sa2))" |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
637 |
and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)" |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
638 |
and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)" |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
639 |
shows "(aa, (t1, t2)) \<approx>gen (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
640 |
(pia + pi) (ac, (sa1, sa2))" |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
641 |
using b c apply - |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
642 |
apply(simp add: alphas2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
643 |
apply(simp add: alphas) |
1581
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
644 |
apply(erule conjE)+ |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
645 |
apply(simp add: fresh_star_plus) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
646 |
apply(drule_tac x="- pia \<bullet> sa1" in spec) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
647 |
apply(drule mp) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
648 |
apply(rotate_tac 5) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
649 |
apply(drule_tac pi="- pia" in r1) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
650 |
apply(simp) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
651 |
apply(rotate_tac -1) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
652 |
apply(drule_tac pi="pia" in r1) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
653 |
apply(simp) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
654 |
apply(drule_tac x="- pia \<bullet> sa2" in spec) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
655 |
apply(drule mp) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
656 |
apply(rotate_tac 6) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
657 |
apply(drule_tac pi="- pia" in r2) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
658 |
apply(simp) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
659 |
apply(rotate_tac -1) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
660 |
apply(drule_tac pi="pia" in r2) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
661 |
apply(simp) |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
662 |
done |
6b1eea8dcdc0
equivp_cheat can be removed for all one-permutation examples.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1563
diff
changeset
|
663 |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
664 |
lemma alpha_res_compose_trans2: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
665 |
fixes pi pia |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
666 |
assumes b: "(aa, (t1, t2)) \<approx>res |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
667 |
(\<lambda>(b, a) (d, c). R1 b d \<and> (\<forall>z. R1 d z \<longrightarrow> R1 b z) \<and> R2 a c \<and> (\<forall>z. R2 c z \<longrightarrow> R2 a z)) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
668 |
(\<lambda>(b, a). fv_a b \<union> fv_b a) pi (ab, (ta1, ta2))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
669 |
and c: "(ab, (ta1, ta2)) \<approx>res (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
670 |
pia (ac, (sa1, sa2))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
671 |
and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
672 |
and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
673 |
shows "(aa, (t1, t2)) \<approx>res (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
674 |
(pia + pi) (ac, (sa1, sa2))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
675 |
using b c apply - |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
676 |
apply(simp add: alphas2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
677 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
678 |
apply(erule conjE)+ |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
679 |
apply(simp add: fresh_star_plus) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
680 |
apply(drule_tac x="- pia \<bullet> sa1" in spec) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
681 |
apply(drule mp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
682 |
apply(rotate_tac 5) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
683 |
apply(drule_tac pi="- pia" in r1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
684 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
685 |
apply(rotate_tac -1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
686 |
apply(drule_tac pi="pia" in r1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
687 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
688 |
apply(drule_tac x="- pia \<bullet> sa2" in spec) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
689 |
apply(drule mp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
690 |
apply(rotate_tac 6) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
691 |
apply(drule_tac pi="- pia" in r2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
692 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
693 |
apply(rotate_tac -1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
694 |
apply(drule_tac pi="pia" in r2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
695 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
696 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
697 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
698 |
lemma alpha_lst_compose_trans2: |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
699 |
fixes pi pia |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
700 |
assumes b: "(aa, (t1, t2)) \<approx>lst |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
701 |
(\<lambda>(b, a) (d, c). R1 b d \<and> (\<forall>z. R1 d z \<longrightarrow> R1 b z) \<and> R2 a c \<and> (\<forall>z. R2 c z \<longrightarrow> R2 a z)) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
702 |
(\<lambda>(b, a). fv_a b \<union> fv_b a) pi (ab, (ta1, ta2))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
703 |
and c: "(ab, (ta1, ta2)) \<approx>lst (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
704 |
pia (ac, (sa1, sa2))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
705 |
and r1: "\<And>pi t s. R1 t s \<Longrightarrow> R1 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
706 |
and r2: "\<And>pi t s. R2 t s \<Longrightarrow> R2 (pi \<bullet> t) (pi \<bullet> s)" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
707 |
shows "(aa, (t1, t2)) \<approx>lst (\<lambda>(b, a) (d, c). R1 b d \<and> R2 a c) (\<lambda>(b, a). fv_a b \<union> fv_b a) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
708 |
(pia + pi) (ac, (sa1, sa2))" |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
709 |
using b c apply - |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
710 |
apply(simp add: alphas2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
711 |
apply(simp add: alphas) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
712 |
apply(erule conjE)+ |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
713 |
apply(simp add: fresh_star_plus) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
714 |
apply(drule_tac x="- pia \<bullet> sa1" in spec) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
715 |
apply(drule mp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
716 |
apply(rotate_tac 5) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
717 |
apply(drule_tac pi="- pia" in r1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
718 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
719 |
apply(rotate_tac -1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
720 |
apply(drule_tac pi="pia" in r1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
721 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
722 |
apply(drule_tac x="- pia \<bullet> sa2" in spec) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
723 |
apply(drule mp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
724 |
apply(rotate_tac 6) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
725 |
apply(drule_tac pi="- pia" in r2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
726 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
727 |
apply(rotate_tac -1) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
728 |
apply(drule_tac pi="pia" in r2) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
729 |
apply(simp) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
730 |
done |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
731 |
|
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
732 |
lemmas alphas_compose_trans2 = alpha_gen_compose_trans2 alpha_res_compose_trans2 alpha_lst_compose_trans2 |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
733 |
|
1657 | 734 |
lemma alpha_gen_refl: |
735 |
assumes a: "R x x" |
|
736 |
shows "(bs, x) \<approx>gen R f 0 (bs, x)" |
|
737 |
and "(bs, x) \<approx>res R f 0 (bs, x)" |
|
738 |
and "(cs, x) \<approx>lst R f 0 (cs, x)" |
|
739 |
using a |
|
740 |
unfolding alphas |
|
741 |
unfolding fresh_star_def |
|
742 |
by (simp_all add: fresh_zero_perm) |
|
743 |
||
744 |
lemma alpha_gen_sym: |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
745 |
assumes a: "R (p \<bullet> x) y \<Longrightarrow> R (- p \<bullet> y) x" |
1657 | 746 |
shows "(bs, x) \<approx>gen R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>gen R f (- p) (bs, x)" |
747 |
and "(bs, x) \<approx>res R f p (cs, y) \<Longrightarrow> (cs, y) \<approx>res R f (- p) (bs, x)" |
|
748 |
and "(ds, x) \<approx>lst R f p (es, y) \<Longrightarrow> (es, y) \<approx>lst R f (- p) (ds, x)" |
|
1664
aa999d263b10
Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1657
diff
changeset
|
749 |
unfolding alphas fresh_star_def |
1657 | 750 |
using a |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
751 |
by (auto simp add: fresh_minus_perm) |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
752 |
|
1657 | 753 |
lemma alpha_gen_trans: |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
754 |
assumes a: "\<lbrakk>R (p \<bullet> x) y; R (q \<bullet> y) z\<rbrakk> \<Longrightarrow> R ((q + p) \<bullet> x) z" |
1657 | 755 |
shows "\<lbrakk>(bs, x) \<approx>gen R f p (cs, y); (cs, y) \<approx>gen R f q (ds, z)\<rbrakk> \<Longrightarrow> (bs, x) \<approx>gen R f (q + p) (ds, z)" |
756 |
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)" |
|
757 |
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)" |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
758 |
using a |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
759 |
unfolding alphas fresh_star_def |
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1666
diff
changeset
|
760 |
by (simp_all add: fresh_plus_perm) |
1657 | 761 |
|
1804
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
762 |
|
2104
2205b572bc9b
Declare alpha_gen_eqvt as eqvt and change the proofs that used 'eqvts[symmetric]'
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2092
diff
changeset
|
763 |
lemma alpha_gen_eqvt[eqvt]: |
1804
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
764 |
shows "(bs, x) \<approx>gen R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>gen (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)" |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
765 |
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)" |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
766 |
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)" |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
767 |
unfolding alphas |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
768 |
unfolding permute_eqvt[symmetric] |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
769 |
unfolding set_eqvt[symmetric] |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
770 |
unfolding permute_fun_app_eq[symmetric] |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
771 |
unfolding Diff_eqvt[symmetric] |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
772 |
by (auto simp add: permute_bool_def fresh_star_permute_iff) |
81b171e2d6d5
added eqvt rules that are more standard
Christian Urban <urbanc@in.tum.de>
parents:
1744
diff
changeset
|
773 |
|
1691
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
774 |
lemma alpha_gen_simpler: |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
775 |
assumes fv_rsp: "\<And>x y. R y x \<Longrightarrow> f x = f y" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
776 |
and fin_fv: "finite (f x)" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
777 |
and fv_eqvt: "pi \<bullet> f x = f (pi \<bullet> x)" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
778 |
shows "alpha_gen (bs, x) R f pi (cs, y) \<longleftrightarrow> |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
779 |
(f x - bs) \<sharp>* pi \<and> |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
780 |
R (pi \<bullet> x) y \<and> |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
781 |
pi \<bullet> bs = cs" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
782 |
apply rule |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
783 |
unfolding alpha_gen |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
784 |
apply clarify |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
785 |
apply (erule conjE)+ |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
786 |
apply (simp) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
787 |
apply (subgoal_tac "f y - cs = pi \<bullet> (f x - bs)") |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
788 |
apply (rule sym) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
789 |
apply simp |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
790 |
apply (rule supp_perm_eq) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
791 |
apply (subst supp_finite_atom_set) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
792 |
apply (rule finite_Diff) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
793 |
apply (rule fin_fv) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
794 |
apply (assumption) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
795 |
apply (simp add: eqvts fv_eqvt) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
796 |
apply (subst fv_rsp) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
797 |
apply assumption |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
798 |
apply (simp) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
799 |
done |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
800 |
|
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
801 |
lemma alpha_lst_simpler: |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
802 |
assumes fv_rsp: "\<And>x y. R y x \<Longrightarrow> f x = f y" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
803 |
and fin_fv: "finite (f x)" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
804 |
and fv_eqvt: "pi \<bullet> f x = f (pi \<bullet> x)" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
805 |
shows "alpha_lst (bs, x) R f pi (cs, y) \<longleftrightarrow> |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
806 |
(f x - set bs) \<sharp>* pi \<and> |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
807 |
R (pi \<bullet> x) y \<and> |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
808 |
pi \<bullet> bs = cs" |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
809 |
apply rule |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
810 |
unfolding alpha_lst |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
811 |
apply clarify |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
812 |
apply (erule conjE)+ |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
813 |
apply (simp) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
814 |
apply (subgoal_tac "f y - set cs = pi \<bullet> (f x - set bs)") |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
815 |
apply (rule sym) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
816 |
apply simp |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
817 |
apply (rule supp_perm_eq) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
818 |
apply (subst supp_finite_atom_set) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
819 |
apply (rule finite_Diff) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
820 |
apply (rule fin_fv) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
821 |
apply (assumption) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
822 |
apply (simp add: eqvts fv_eqvt) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
823 |
apply (subst fv_rsp) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
824 |
apply assumption |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
825 |
apply (simp) |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
826 |
done |
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
827 |
|
2068
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
828 |
fun |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
829 |
prod_fv :: "('a \<Rightarrow> atom set) \<Rightarrow> ('b \<Rightarrow> atom set) \<Rightarrow> ('a \<times> 'b) \<Rightarrow> atom set" |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
830 |
where |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
831 |
"prod_fv fvx fvy (x, y) = (fvx x \<union> fvy y)" |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
832 |
|
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
833 |
lemma [quot_respect]: |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
834 |
"((R1 ===> op =) ===> (R2 ===> op =) ===> prod_rel R1 R2 ===> op =) prod_fv prod_fv" |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
835 |
by auto |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
836 |
|
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
837 |
lemma [quot_preserve]: |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
838 |
assumes q1: "Quotient R1 abs1 rep1" |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
839 |
and q2: "Quotient R2 abs2 rep2" |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
840 |
shows "((abs1 ---> id) ---> (abs2 ---> id) ---> prod_fun rep1 rep2 ---> id) prod_fv = prod_fv" |
79b733010bc5
prod_fv and its respectfullness and preservation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1933
diff
changeset
|
841 |
by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) |
1691
b497ac81aead
Abs_gen and Abs_let simplifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1688
diff
changeset
|
842 |
|
2092
c0ab7451b20d
prod_rel and prod_fv eqvt and mono
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2068
diff
changeset
|
843 |
lemma [mono]: "A <= B \<Longrightarrow> C <= D ==> prod_rel A C <= prod_rel B D" |
c0ab7451b20d
prod_rel and prod_fv eqvt and mono
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2068
diff
changeset
|
844 |
by auto |
c0ab7451b20d
prod_rel and prod_fv eqvt and mono
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2068
diff
changeset
|
845 |
|
2124
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
846 |
lemma [eqvt]: |
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
847 |
shows "p \<bullet> prod_rel A B x y = prod_rel (p \<bullet> A) (p \<bullet> B) (p \<bullet> x) (p \<bullet> y)" |
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
848 |
unfolding prod_rel.simps |
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
849 |
by (perm_simp) (rule refl) |
2092
c0ab7451b20d
prod_rel and prod_fv eqvt and mono
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2068
diff
changeset
|
850 |
|
2124
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
851 |
lemma [eqvt]: |
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
852 |
shows "p \<bullet> prod_fv A B (x, y) = prod_fv (p \<bullet> A) (p \<bullet> B) (p \<bullet> x, p \<bullet> y)" |
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
853 |
unfolding prod_fv.simps |
a17b25cb94a6
tuned eqvt-proofs about prod_rel and prod_fv
Christian Urban <urbanc@in.tum.de>
parents:
2104
diff
changeset
|
854 |
by (perm_simp) (rule refl) |
2092
c0ab7451b20d
prod_rel and prod_fv eqvt and mono
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2068
diff
changeset
|
855 |
|
1440
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
856 |
end |
ffd5540ac2e9
added preliminary test version....but Test works now
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
857 |