58 done |
59 done |
59 |
60 |
60 end |
61 end |
61 |
62 |
62 (* |
63 (* |
63 setup {* snd o define_raw_perms ["kind", "ty", "trm"] ["LFex.kind", "LFex.ty", "LFex.trm"] *} |
64 setup {* snd o define_raw_perms ["rkind", "rty", "rtrm"] ["LFex.rkind", "LFex.rty", "LFex.rtrm"] *} |
64 local_setup {* |
65 local_setup {* |
65 snd o define_fv_alpha "LFex.kind" |
66 snd o define_fv_alpha "LFex.rkind" |
66 [[ [], [[], [(NONE, 1)], [(NONE, 1)]] ], |
67 [[ [], [[], [(NONE, 1)], [(NONE, 1)]] ], |
67 [ [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]] ], |
68 [ [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]] ], |
68 [ [[]], [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]]]] *} |
69 [ [[]], [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]]]] *} |
69 notation |
70 notation |
70 alpha_kind ("_ \<approx>ki _" [100, 100] 100) |
71 alpha_rkind ("_ \<approx>ki _" [100, 100] 100) |
71 and alpha_ty ("_ \<approx>ty _" [100, 100] 100) |
72 and alpha_rty ("_ \<approx>rty _" [100, 100] 100) |
72 and alpha_trm ("_ \<approx>tr _" [100, 100] 100) |
73 and alpha_rtrm ("_ \<approx>tr _" [100, 100] 100) |
73 thm fv_kind_fv_ty_fv_trm.simps alpha_kind_alpha_ty_alpha_trm.intros |
74 thm fv_rkind_fv_rty_fv_rtrm.simps alpha_rkind_alpha_rty_alpha_rtrm.intros |
74 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alphas_inj}, []), (build_alpha_inj @{thms alpha_kind_alpha_ty_alpha_trm.intros} @{thms kind.distinct ty.distinct trm.distinct kind.inject ty.inject trm.inject} @{thms alpha_kind.cases alpha_ty.cases alpha_trm.cases} ctxt)) ctxt)) *} |
75 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alphas_inj}, []), (build_alpha_inj @{thms alpha_rkind_alpha_rty_alpha_rtrm.intros} @{thms rkind.distinct rty.distinct rtrm.distinct rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases} ctxt)) ctxt)) *} |
75 thm alphas_inj |
76 thm alphas_inj |
76 |
77 |
77 lemma alphas_eqvt: |
78 lemma alphas_eqvt: |
78 "t1 \<approx>ki s1 \<Longrightarrow> (pi \<bullet> t1) \<approx>ki (pi \<bullet> s1)" |
79 "t1 \<approx>ki s1 \<Longrightarrow> (pi \<bullet> t1) \<approx>ki (pi \<bullet> s1)" |
79 "t2 \<approx>ty s2 \<Longrightarrow> (pi \<bullet> t2) \<approx>ty (pi \<bullet> s2)" |
80 "t2 \<approx>rty s2 \<Longrightarrow> (pi \<bullet> t2) \<approx>rty (pi \<bullet> s2)" |
80 "t3 \<approx>tr s3 \<Longrightarrow> (pi \<bullet> t3) \<approx>tr (pi \<bullet> s3)" |
81 "t3 \<approx>tr s3 \<Longrightarrow> (pi \<bullet> t3) \<approx>tr (pi \<bullet> s3)" |
81 sorry |
82 sorry |
82 |
83 |
83 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alphas_equivp}, []), |
84 local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alphas_equivp}, []), |
84 (build_equivps [@{term alpha_kind}, @{term alpha_ty}, @{term alpha_trm}] |
85 (build_equivps [@{term alpha_rkind}, @{term alpha_rty}, @{term alpha_rtrm}] |
85 @{thm kind_ty_trm.induct} @{thm alpha_kind_alpha_ty_alpha_trm.induct} |
86 @{thm rkind_rty_rtrm.induct} @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct} |
86 @{thms kind.inject ty.inject trm.inject} @{thms alphas_inj} |
87 @{thms rkind.inject rty.inject rtrm.inject} @{thms alphas_inj} |
87 @{thms kind.distinct ty.distinct trm.distinct} |
88 @{thms rkind.distinct rty.distinct rtrm.distinct} |
88 @{thms alpha_kind.cases alpha_ty.cases alpha_trm.cases} |
89 @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases} |
89 @{thms alphas_eqvt} ctxt)) ctxt)) *} |
90 @{thms alphas_eqvt} ctxt)) ctxt)) *} |
90 thm alphas_equivp |
91 thm alphas_equivp |
91 *) |
92 *) |
92 |
93 |
93 primrec |
94 primrec |
94 rfv_kind :: "kind \<Rightarrow> atom set" |
95 rfv_rkind :: "rkind \<Rightarrow> atom set" |
95 and rfv_ty :: "ty \<Rightarrow> atom set" |
96 and rfv_rty :: "rty \<Rightarrow> atom set" |
96 and rfv_trm :: "trm \<Rightarrow> atom set" |
97 and rfv_rtrm :: "rtrm \<Rightarrow> atom set" |
97 where |
98 where |
98 "rfv_kind (Type) = {}" |
99 "rfv_rkind (Type) = {}" |
99 | "rfv_kind (KPi A x K) = (rfv_ty A) \<union> ((rfv_kind K) - {atom x})" |
100 | "rfv_rkind (KPi A x K) = (rfv_rty A) \<union> ((rfv_rkind K) - {atom x})" |
100 | "rfv_ty (TConst i) = {atom i}" |
101 | "rfv_rty (TConst i) = {atom i}" |
101 | "rfv_ty (TApp A M) = (rfv_ty A) \<union> (rfv_trm M)" |
102 | "rfv_rty (TApp A M) = (rfv_rty A) \<union> (rfv_rtrm M)" |
102 | "rfv_ty (TPi A x B) = (rfv_ty A) \<union> ((rfv_ty B) - {atom x})" |
103 | "rfv_rty (TPi A x B) = (rfv_rty A) \<union> ((rfv_rty B) - {atom x})" |
103 | "rfv_trm (Const i) = {atom i}" |
104 | "rfv_rtrm (Const i) = {atom i}" |
104 | "rfv_trm (Var x) = {atom x}" |
105 | "rfv_rtrm (Var x) = {atom x}" |
105 | "rfv_trm (App M N) = (rfv_trm M) \<union> (rfv_trm N)" |
106 | "rfv_rtrm (App M N) = (rfv_rtrm M) \<union> (rfv_rtrm N)" |
106 | "rfv_trm (Lam A x M) = (rfv_ty A) \<union> ((rfv_trm M) - {atom x})" |
107 | "rfv_rtrm (Lam A x M) = (rfv_rty A) \<union> ((rfv_rtrm M) - {atom x})" |
107 |
108 |
108 inductive |
109 inductive |
109 akind :: "kind \<Rightarrow> kind \<Rightarrow> bool" ("_ \<approx>ki _" [100, 100] 100) |
110 arkind :: "rkind \<Rightarrow> rkind \<Rightarrow> bool" ("_ \<approx>ki _" [100, 100] 100) |
110 and aty :: "ty \<Rightarrow> ty \<Rightarrow> bool" ("_ \<approx>ty _" [100, 100] 100) |
111 and arty :: "rty \<Rightarrow> rty \<Rightarrow> bool" ("_ \<approx>rty _" [100, 100] 100) |
111 and atrm :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<approx>tr _" [100, 100] 100) |
112 and artrm :: "rtrm \<Rightarrow> rtrm \<Rightarrow> bool" ("_ \<approx>tr _" [100, 100] 100) |
112 where |
113 where |
113 a1: "(Type) \<approx>ki (Type)" |
114 a1: "(Type) \<approx>ki (Type)" |
114 | a2: "\<lbrakk>A \<approx>ty A'; \<exists>pi. (({atom a}, K) \<approx>gen akind rfv_kind pi ({atom b}, K'))\<rbrakk> \<Longrightarrow> (KPi A a K) \<approx>ki (KPi A' b K')" |
115 | a2: "\<lbrakk>A \<approx>rty A'; \<exists>pi. (({atom a}, K) \<approx>gen arkind rfv_rkind pi ({atom b}, K'))\<rbrakk> \<Longrightarrow> (KPi A a K) \<approx>ki (KPi A' b K')" |
115 | a3: "i = j \<Longrightarrow> (TConst i) \<approx>ty (TConst j)" |
116 | a3: "i = j \<Longrightarrow> (TConst i) \<approx>rty (TConst j)" |
116 | a4: "\<lbrakk>A \<approx>ty A'; M \<approx>tr M'\<rbrakk> \<Longrightarrow> (TApp A M) \<approx>ty (TApp A' M')" |
117 | a4: "\<lbrakk>A \<approx>rty A'; M \<approx>tr M'\<rbrakk> \<Longrightarrow> (TApp A M) \<approx>rty (TApp A' M')" |
117 | a5: "\<lbrakk>A \<approx>ty A'; \<exists>pi. (({atom a}, B) \<approx>gen aty rfv_ty pi ({atom b}, B'))\<rbrakk> \<Longrightarrow> (TPi A a B) \<approx>ty (TPi A' b B')" |
118 | a5: "\<lbrakk>A \<approx>rty A'; \<exists>pi. (({atom a}, B) \<approx>gen arty rfv_rty pi ({atom b}, B'))\<rbrakk> \<Longrightarrow> (TPi A a B) \<approx>rty (TPi A' b B')" |
118 | a6: "i = j \<Longrightarrow> (Const i) \<approx>tr (Const j)" |
119 | a6: "i = j \<Longrightarrow> (Const i) \<approx>tr (Const j)" |
119 | a7: "x = y \<Longrightarrow> (Var x) \<approx>tr (Var y)" |
120 | a7: "x = y \<Longrightarrow> (Var x) \<approx>tr (Var y)" |
120 | a8: "\<lbrakk>M \<approx>tr M'; N \<approx>tr N'\<rbrakk> \<Longrightarrow> (App M N) \<approx>tr (App M' N')" |
121 | a8: "\<lbrakk>M \<approx>tr M'; N \<approx>tr N'\<rbrakk> \<Longrightarrow> (App M N) \<approx>tr (App M' N')" |
121 | a9: "\<lbrakk>A \<approx>ty A'; \<exists>pi. (({atom a}, M) \<approx>gen atrm rfv_trm pi ({atom b}, M'))\<rbrakk> \<Longrightarrow> (Lam A a M) \<approx>tr (Lam A' b M')" |
122 | a9: "\<lbrakk>A \<approx>rty A'; \<exists>pi. (({atom a}, M) \<approx>gen artrm rfv_rtrm pi ({atom b}, M'))\<rbrakk> \<Longrightarrow> (Lam A a M) \<approx>tr (Lam A' b M')" |
122 |
123 |
123 lemma akind_aty_atrm_inj: |
124 lemma arkind_arty_artrm_inj: |
124 "(Type) \<approx>ki (Type) = True" |
125 "(Type) \<approx>ki (Type) = True" |
125 "((KPi A a K) \<approx>ki (KPi A' b K')) = ((A \<approx>ty A') \<and> (\<exists>pi. ({atom a}, K) \<approx>gen akind rfv_kind pi ({atom b}, K')))" |
126 "((KPi A a K) \<approx>ki (KPi A' b K')) = ((A \<approx>rty A') \<and> (\<exists>pi. ({atom a}, K) \<approx>gen arkind rfv_rkind pi ({atom b}, K')))" |
126 "(TConst i) \<approx>ty (TConst j) = (i = j)" |
127 "(TConst i) \<approx>rty (TConst j) = (i = j)" |
127 "(TApp A M) \<approx>ty (TApp A' M') = (A \<approx>ty A' \<and> M \<approx>tr M')" |
128 "(TApp A M) \<approx>rty (TApp A' M') = (A \<approx>rty A' \<and> M \<approx>tr M')" |
128 "((TPi A a B) \<approx>ty (TPi A' b B')) = ((A \<approx>ty A') \<and> (\<exists>pi. (({atom a}, B) \<approx>gen aty rfv_ty pi ({atom b}, B'))))" |
129 "((TPi A a B) \<approx>rty (TPi A' b B')) = ((A \<approx>rty A') \<and> (\<exists>pi. (({atom a}, B) \<approx>gen arty rfv_rty pi ({atom b}, B'))))" |
129 "(Const i) \<approx>tr (Const j) = (i = j)" |
130 "(Const i) \<approx>tr (Const j) = (i = j)" |
130 "(Var x) \<approx>tr (Var y) = (x = y)" |
131 "(Var x) \<approx>tr (Var y) = (x = y)" |
131 "(App M N) \<approx>tr (App M' N') = (M \<approx>tr M' \<and> N \<approx>tr N')" |
132 "(App M N) \<approx>tr (App M' N') = (M \<approx>tr M' \<and> N \<approx>tr N')" |
132 "((Lam A a M) \<approx>tr (Lam A' b M')) = ((A \<approx>ty A') \<and> (\<exists>pi. (({atom a}, M) \<approx>gen atrm rfv_trm pi ({atom b}, M'))))" |
133 "((Lam A a M) \<approx>tr (Lam A' b M')) = ((A \<approx>rty A') \<and> (\<exists>pi. (({atom a}, M) \<approx>gen artrm rfv_rtrm pi ({atom b}, M'))))" |
133 apply - |
134 apply - |
134 apply (simp add: akind_aty_atrm.intros) |
135 apply (simp add: arkind_arty_artrm.intros) |
135 |
136 |
136 apply rule apply (erule akind.cases) apply simp apply blast |
137 apply rule apply (erule arkind.cases) apply simp apply blast |
137 apply (simp only: akind_aty_atrm.intros) |
138 apply (simp only: arkind_arty_artrm.intros) |
138 |
139 |
139 apply rule apply (erule aty.cases) apply simp apply simp apply simp |
140 apply rule apply (erule arty.cases) apply simp apply simp apply simp |
140 apply (simp only: akind_aty_atrm.intros) |
141 apply (simp only: arkind_arty_artrm.intros) |
141 |
142 |
142 apply rule apply (erule aty.cases) apply simp apply simp apply simp |
143 apply rule apply (erule arty.cases) apply simp apply simp apply simp |
143 apply (simp only: akind_aty_atrm.intros) |
144 apply (simp only: arkind_arty_artrm.intros) |
144 |
145 |
145 apply rule apply (erule aty.cases) apply simp apply simp apply blast |
146 apply rule apply (erule arty.cases) apply simp apply simp apply blast |
146 apply (simp only: akind_aty_atrm.intros) |
147 apply (simp only: arkind_arty_artrm.intros) |
147 |
148 |
148 apply rule apply (erule atrm.cases) apply simp apply simp apply simp apply blast |
149 apply rule apply (erule artrm.cases) apply simp apply simp apply simp apply blast |
149 apply (simp only: akind_aty_atrm.intros) |
150 apply (simp only: arkind_arty_artrm.intros) |
150 |
151 |
151 apply rule apply (erule atrm.cases) apply simp apply simp apply simp apply blast |
152 apply rule apply (erule artrm.cases) apply simp apply simp apply simp apply blast |
152 apply (simp only: akind_aty_atrm.intros) |
153 apply (simp only: arkind_arty_artrm.intros) |
153 |
154 |
154 apply rule apply (erule atrm.cases) apply simp apply simp apply simp apply blast |
155 apply rule apply (erule artrm.cases) apply simp apply simp apply simp apply blast |
155 apply (simp only: akind_aty_atrm.intros) |
156 apply (simp only: arkind_arty_artrm.intros) |
156 |
157 |
157 apply rule apply (erule atrm.cases) apply simp apply simp apply simp apply blast |
158 apply rule apply (erule artrm.cases) apply simp apply simp apply simp apply blast |
158 apply (simp only: akind_aty_atrm.intros) |
159 apply (simp only: arkind_arty_artrm.intros) |
159 done |
160 done |
160 |
161 |
161 lemma rfv_eqvt[eqvt]: |
162 lemma rfv_eqvt[eqvt]: |
162 "((pi\<bullet>rfv_kind t1) = rfv_kind (pi\<bullet>t1))" |
163 "((pi\<bullet>rfv_rkind t1) = rfv_rkind (pi\<bullet>t1))" |
163 "((pi\<bullet>rfv_ty t2) = rfv_ty (pi\<bullet>t2))" |
164 "((pi\<bullet>rfv_rty t2) = rfv_rty (pi\<bullet>t2))" |
164 "((pi\<bullet>rfv_trm t3) = rfv_trm (pi\<bullet>t3))" |
165 "((pi\<bullet>rfv_rtrm t3) = rfv_rtrm (pi\<bullet>t3))" |
165 apply(induct t1 and t2 and t3 rule: kind_ty_trm.inducts) |
166 apply(induct t1 and t2 and t3 rule: rkind_rty_rtrm.inducts) |
166 apply(simp_all add: union_eqvt Diff_eqvt) |
167 apply(simp_all add: union_eqvt Diff_eqvt) |
167 apply(simp_all add: permute_set_eq atom_eqvt) |
168 apply(simp_all add: permute_set_eq atom_eqvt) |
168 done |
169 done |
169 |
170 |
170 lemma alpha_eqvt: |
171 lemma alpha_eqvt: |
171 "t1 \<approx>ki s1 \<Longrightarrow> (pi \<bullet> t1) \<approx>ki (pi \<bullet> s1)" |
172 "t1 \<approx>ki s1 \<Longrightarrow> (pi \<bullet> t1) \<approx>ki (pi \<bullet> s1)" |
172 "t2 \<approx>ty s2 \<Longrightarrow> (pi \<bullet> t2) \<approx>ty (pi \<bullet> s2)" |
173 "t2 \<approx>rty s2 \<Longrightarrow> (pi \<bullet> t2) \<approx>rty (pi \<bullet> s2)" |
173 "t3 \<approx>tr s3 \<Longrightarrow> (pi \<bullet> t3) \<approx>tr (pi \<bullet> s3)" |
174 "t3 \<approx>tr s3 \<Longrightarrow> (pi \<bullet> t3) \<approx>tr (pi \<bullet> s3)" |
174 apply(induct rule: akind_aty_atrm.inducts) |
175 apply(induct rule: arkind_arty_artrm.inducts) |
175 apply (simp_all add: akind_aty_atrm.intros) |
176 apply (simp_all add: arkind_arty_artrm.intros) |
176 apply (simp_all add: akind_aty_atrm_inj) |
177 apply (simp_all add: arkind_arty_artrm_inj) |
177 apply (rule alpha_gen_atom_eqvt) |
178 apply (rule alpha_gen_atom_eqvt) |
178 apply (simp add: rfv_eqvt) |
179 apply (simp add: rfv_eqvt) |
179 apply assumption |
180 apply assumption |
180 apply (rule alpha_gen_atom_eqvt) |
181 apply (rule alpha_gen_atom_eqvt) |
181 apply (simp add: rfv_eqvt) |
182 apply (simp add: rfv_eqvt) |
207 apply(rule_tac x="0" in exI) |
208 apply(rule_tac x="0" in exI) |
208 apply(simp_all add: fresh_star_def fresh_zero_perm alpha_gen) |
209 apply(simp_all add: fresh_star_def fresh_zero_perm alpha_gen) |
209 done |
210 done |
210 |
211 |
211 lemma al_sym: |
212 lemma al_sym: |
212 fixes K K'::"kind" and A A'::"ty" and M M'::"trm" |
213 fixes K K'::"rkind" and A A'::"rty" and M M'::"rtrm" |
213 shows "K \<approx>ki K' \<Longrightarrow> K' \<approx>ki K" |
214 shows "K \<approx>ki K' \<Longrightarrow> K' \<approx>ki K" |
214 and "A \<approx>ty A' \<Longrightarrow> A' \<approx>ty A" |
215 and "A \<approx>rty A' \<Longrightarrow> A' \<approx>rty A" |
215 and "M \<approx>tr M' \<Longrightarrow> M' \<approx>tr M" |
216 and "M \<approx>tr M' \<Longrightarrow> M' \<approx>tr M" |
216 apply(induct K K' and A A' and M M' rule: akind_aty_atrm.inducts) |
217 apply(induct K K' and A A' and M M' rule: arkind_arty_artrm.inducts) |
217 apply(simp_all add: akind_aty_atrm.intros) |
218 apply(simp_all add: arkind_arty_artrm.intros) |
218 apply (simp_all add: akind_aty_atrm_inj) |
219 apply (simp_all add: arkind_arty_artrm_inj) |
219 apply(erule alpha_gen_compose_sym) |
220 apply(erule alpha_gen_compose_sym) |
220 apply(erule alpha_eqvt) |
221 apply(erule alpha_eqvt) |
221 apply(erule alpha_gen_compose_sym) |
222 apply(erule alpha_gen_compose_sym) |
222 apply(erule alpha_eqvt) |
223 apply(erule alpha_eqvt) |
223 apply(erule alpha_gen_compose_sym) |
224 apply(erule alpha_gen_compose_sym) |
224 apply(erule alpha_eqvt) |
225 apply(erule alpha_eqvt) |
225 done |
226 done |
226 |
227 |
227 lemma al_trans: |
228 lemma al_trans: |
228 fixes K K' K''::"kind" and A A' A''::"ty" and M M' M''::"trm" |
229 fixes K K' K''::"rkind" and A A' A''::"rty" and M M' M''::"rtrm" |
229 shows "K \<approx>ki K' \<Longrightarrow> K' \<approx>ki K'' \<Longrightarrow> K \<approx>ki K''" |
230 shows "K \<approx>ki K' \<Longrightarrow> K' \<approx>ki K'' \<Longrightarrow> K \<approx>ki K''" |
230 and "A \<approx>ty A' \<Longrightarrow> A' \<approx>ty A'' \<Longrightarrow> A \<approx>ty A''" |
231 and "A \<approx>rty A' \<Longrightarrow> A' \<approx>rty A'' \<Longrightarrow> A \<approx>rty A''" |
231 and "M \<approx>tr M' \<Longrightarrow> M' \<approx>tr M'' \<Longrightarrow> M \<approx>tr M''" |
232 and "M \<approx>tr M' \<Longrightarrow> M' \<approx>tr M'' \<Longrightarrow> M \<approx>tr M''" |
232 apply(induct K K' and A A' and M M' arbitrary: K'' and A'' and M'' rule: akind_aty_atrm.inducts) |
233 apply(induct K K' and A A' and M M' arbitrary: K'' and A'' and M'' rule: arkind_arty_artrm.inducts) |
233 apply(simp_all add: akind_aty_atrm.intros) |
234 apply(simp_all add: arkind_arty_artrm.intros) |
234 apply(erule akind.cases) |
235 apply(erule arkind.cases) |
235 apply(simp_all add: akind_aty_atrm.intros) |
236 apply(simp_all add: arkind_arty_artrm.intros) |
236 apply(simp add: akind_aty_atrm_inj) |
237 apply(simp add: arkind_arty_artrm_inj) |
237 apply(erule alpha_gen_compose_trans) |
238 apply(erule alpha_gen_compose_trans) |
238 apply(assumption) |
239 apply(assumption) |
239 apply(erule alpha_eqvt) |
240 apply(erule alpha_eqvt) |
240 apply(rotate_tac 4) |
241 apply(rotate_tac 4) |
241 apply(erule aty.cases) |
242 apply(erule arty.cases) |
242 apply(simp_all add: akind_aty_atrm.intros) |
243 apply(simp_all add: arkind_arty_artrm.intros) |
243 apply(rotate_tac 3) |
244 apply(rotate_tac 3) |
244 apply(erule aty.cases) |
245 apply(erule arty.cases) |
245 apply(simp_all add: akind_aty_atrm.intros) |
246 apply(simp_all add: arkind_arty_artrm.intros) |
246 apply(simp add: akind_aty_atrm_inj) |
247 apply(simp add: arkind_arty_artrm_inj) |
247 apply(erule alpha_gen_compose_trans) |
248 apply(erule alpha_gen_compose_trans) |
248 apply(assumption) |
249 apply(assumption) |
249 apply(erule alpha_eqvt) |
250 apply(erule alpha_eqvt) |
250 apply(rotate_tac 4) |
251 apply(rotate_tac 4) |
251 apply(erule atrm.cases) |
252 apply(erule artrm.cases) |
252 apply(simp_all add: akind_aty_atrm.intros) |
253 apply(simp_all add: arkind_arty_artrm.intros) |
253 apply(rotate_tac 3) |
254 apply(rotate_tac 3) |
254 apply(erule atrm.cases) |
255 apply(erule artrm.cases) |
255 apply(simp_all add: akind_aty_atrm.intros) |
256 apply(simp_all add: arkind_arty_artrm.intros) |
256 apply(simp add: akind_aty_atrm_inj) |
257 apply(simp add: arkind_arty_artrm_inj) |
257 apply(erule alpha_gen_compose_trans) |
258 apply(erule alpha_gen_compose_trans) |
258 apply(assumption) |
259 apply(assumption) |
259 apply(erule alpha_eqvt) |
260 apply(erule alpha_eqvt) |
260 done |
261 done |
261 |
262 |
262 lemma alpha_equivps: |
263 lemma alpha_equivps: |
263 shows "equivp akind" |
264 shows "equivp arkind" |
264 and "equivp aty" |
265 and "equivp arty" |
265 and "equivp atrm" |
266 and "equivp artrm" |
266 apply(rule equivpI) |
267 apply(rule equivpI) |
267 unfolding reflp_def symp_def transp_def |
268 unfolding reflp_def symp_def transp_def |
268 apply(auto intro: al_refl al_sym al_trans) |
269 apply(auto intro: al_refl al_sym al_trans) |
269 apply(rule equivpI) |
270 apply(rule equivpI) |
270 unfolding reflp_def symp_def transp_def |
271 unfolding reflp_def symp_def transp_def |
272 apply(rule equivpI) |
273 apply(rule equivpI) |
273 unfolding reflp_def symp_def transp_def |
274 unfolding reflp_def symp_def transp_def |
274 apply(auto intro: al_refl al_sym al_trans) |
275 apply(auto intro: al_refl al_sym al_trans) |
275 done |
276 done |
276 |
277 |
277 quotient_type KIND = kind / akind |
278 quotient_type RKIND = rkind / arkind |
278 by (rule alpha_equivps) |
279 by (rule alpha_equivps) |
279 |
280 |
280 quotient_type |
281 quotient_type |
281 TY = ty / aty and |
282 RTY = rty / arty and |
282 TRM = trm / atrm |
283 RTRM = rtrm / artrm |
283 by (auto intro: alpha_equivps) |
284 by (auto intro: alpha_equivps) |
284 |
285 |
285 quotient_definition |
286 quotient_definition |
286 "TYP :: KIND" |
287 "TYP :: RKIND" |
287 is |
288 is |
288 "Type" |
289 "Type" |
289 |
290 |
290 quotient_definition |
291 quotient_definition |
291 "KPI :: TY \<Rightarrow> name \<Rightarrow> KIND \<Rightarrow> KIND" |
292 "KPI :: RTY \<Rightarrow> name \<Rightarrow> RKIND \<Rightarrow> RKIND" |
292 is |
293 is |
293 "KPi" |
294 "KPi" |
294 |
295 |
295 quotient_definition |
296 quotient_definition |
296 "TCONST :: ident \<Rightarrow> TY" |
297 "TCONST :: ident \<Rightarrow> RTY" |
297 is |
298 is |
298 "TConst" |
299 "TConst" |
299 |
300 |
300 quotient_definition |
301 quotient_definition |
301 "TAPP :: TY \<Rightarrow> TRM \<Rightarrow> TY" |
302 "TAPP :: RTY \<Rightarrow> RTRM \<Rightarrow> RTY" |
302 is |
303 is |
303 "TApp" |
304 "TApp" |
304 |
305 |
305 quotient_definition |
306 quotient_definition |
306 "TPI :: TY \<Rightarrow> name \<Rightarrow> TY \<Rightarrow> TY" |
307 "TPI :: RTY \<Rightarrow> name \<Rightarrow> RTY \<Rightarrow> RTY" |
307 is |
308 is |
308 "TPi" |
309 "TPi" |
309 |
310 |
310 (* FIXME: does not work with CONST *) |
311 (* FIXME: does not work with CONST *) |
311 quotient_definition |
312 quotient_definition |
312 "CONS :: ident \<Rightarrow> TRM" |
313 "CONS :: ident \<Rightarrow> RTRM" |
313 is |
314 is |
314 "Const" |
315 "Const" |
315 |
316 |
316 quotient_definition |
317 quotient_definition |
317 "VAR :: name \<Rightarrow> TRM" |
318 "VAR :: name \<Rightarrow> RTRM" |
318 is |
319 is |
319 "Var" |
320 "Var" |
320 |
321 |
321 quotient_definition |
322 quotient_definition |
322 "APP :: TRM \<Rightarrow> TRM \<Rightarrow> TRM" |
323 "APP :: RTRM \<Rightarrow> RTRM \<Rightarrow> RTRM" |
323 is |
324 is |
324 "App" |
325 "App" |
325 |
326 |
326 quotient_definition |
327 quotient_definition |
327 "LAM :: TY \<Rightarrow> name \<Rightarrow> TRM \<Rightarrow> TRM" |
328 "LAM :: RTY \<Rightarrow> name \<Rightarrow> RTRM \<Rightarrow> RTRM" |
328 is |
329 is |
329 "Lam" |
330 "Lam" |
330 |
331 |
331 (* FIXME: print out a warning if the type contains a liftet type, like kind \<Rightarrow> name set *) |
332 (* FIXME: print out a warning if the type contains a liftet type, like rkind \<Rightarrow> name set *) |
332 quotient_definition |
333 quotient_definition |
333 "fv_kind :: KIND \<Rightarrow> atom set" |
334 "fv_rkind :: RKIND \<Rightarrow> atom set" |
334 is |
335 is |
335 "rfv_kind" |
336 "rfv_rkind" |
336 |
337 |
337 quotient_definition |
338 quotient_definition |
338 "fv_ty :: TY \<Rightarrow> atom set" |
339 "fv_rty :: RTY \<Rightarrow> atom set" |
339 is |
340 is |
340 "rfv_ty" |
341 "rfv_rty" |
341 |
342 |
342 quotient_definition |
343 quotient_definition |
343 "fv_trm :: TRM \<Rightarrow> atom set" |
344 "fv_rtrm :: RTRM \<Rightarrow> atom set" |
344 is |
345 is |
345 "rfv_trm" |
346 "rfv_rtrm" |
346 |
347 |
347 lemma alpha_rfv: |
348 lemma alpha_rfv: |
348 shows "(t \<approx>ki s \<longrightarrow> rfv_kind t = rfv_kind s) \<and> |
349 shows "(t \<approx>ki s \<longrightarrow> rfv_rkind t = rfv_rkind s) \<and> |
349 (t1 \<approx>ty s1 \<longrightarrow> rfv_ty t1 = rfv_ty s1) \<and> |
350 (t1 \<approx>rty s1 \<longrightarrow> rfv_rty t1 = rfv_rty s1) \<and> |
350 (t2 \<approx>tr s2 \<longrightarrow> rfv_trm t2 = rfv_trm s2)" |
351 (t2 \<approx>tr s2 \<longrightarrow> rfv_rtrm t2 = rfv_rtrm s2)" |
351 apply(rule akind_aty_atrm.induct) |
352 apply(rule arkind_arty_artrm.induct) |
352 apply(simp_all add: alpha_gen) |
353 apply(simp_all add: alpha_gen) |
353 done |
354 done |
354 |
355 |
355 lemma perm_rsp[quot_respect]: |
356 lemma perm_rsp[quot_respect]: |
356 "(op = ===> akind ===> akind) permute permute" |
357 "(op = ===> arkind ===> arkind) permute permute" |
357 "(op = ===> aty ===> aty) permute permute" |
358 "(op = ===> arty ===> arty) permute permute" |
358 "(op = ===> atrm ===> atrm) permute permute" |
359 "(op = ===> artrm ===> artrm) permute permute" |
359 by (simp_all add:alpha_eqvt) |
360 by (simp_all add:alpha_eqvt) |
360 |
361 |
361 lemma tconst_rsp[quot_respect]: |
362 lemma tconst_rsp[quot_respect]: |
362 "(op = ===> aty) TConst TConst" |
363 "(op = ===> arty) TConst TConst" |
363 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
364 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
364 lemma tapp_rsp[quot_respect]: |
365 lemma tapp_rsp[quot_respect]: |
365 "(aty ===> atrm ===> aty) TApp TApp" |
366 "(arty ===> artrm ===> arty) TApp TApp" |
366 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
367 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
367 lemma var_rsp[quot_respect]: |
368 lemma var_rsp[quot_respect]: |
368 "(op = ===> atrm) Var Var" |
369 "(op = ===> artrm) Var Var" |
369 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
370 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
370 lemma app_rsp[quot_respect]: |
371 lemma app_rsp[quot_respect]: |
371 "(atrm ===> atrm ===> atrm) App App" |
372 "(artrm ===> artrm ===> artrm) App App" |
372 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
373 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
373 lemma const_rsp[quot_respect]: |
374 lemma const_rsp[quot_respect]: |
374 "(op = ===> atrm) Const Const" |
375 "(op = ===> artrm) Const Const" |
375 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
376 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) done |
376 |
377 |
377 lemma kpi_rsp[quot_respect]: |
378 lemma kpi_rsp[quot_respect]: |
378 "(aty ===> op = ===> akind ===> akind) KPi KPi" |
379 "(arty ===> op = ===> arkind ===> arkind) KPi KPi" |
379 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) |
380 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) |
380 apply (rule a2) apply assumption |
381 apply (rule a2) apply assumption |
381 apply (rule_tac x="0" in exI) |
382 apply (rule_tac x="0" in exI) |
382 apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen) |
383 apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen) |
383 done |
384 done |
384 |
385 |
385 lemma tpi_rsp[quot_respect]: |
386 lemma tpi_rsp[quot_respect]: |
386 "(aty ===> op = ===> aty ===> aty) TPi TPi" |
387 "(arty ===> op = ===> arty ===> arty) TPi TPi" |
387 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) |
388 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) |
388 apply (rule a5) apply assumption |
389 apply (rule a5) apply assumption |
389 apply (rule_tac x="0" in exI) |
390 apply (rule_tac x="0" in exI) |
390 apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen) |
391 apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen) |
391 done |
392 done |
392 lemma lam_rsp[quot_respect]: |
393 lemma lam_rsp[quot_respect]: |
393 "(aty ===> op = ===> atrm ===> atrm) Lam Lam" |
394 "(arty ===> op = ===> artrm ===> artrm) Lam Lam" |
394 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) |
395 apply (auto intro: a1 a2 a3 a4 a5 a6 a7 a8 a9) |
395 apply (rule a9) apply assumption |
396 apply (rule a9) apply assumption |
396 apply (rule_tac x="0" in exI) |
397 apply (rule_tac x="0" in exI) |
397 apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen) |
398 apply (simp add: fresh_star_def fresh_zero_perm alpha_rfv alpha_gen) |
398 done |
399 done |
399 |
400 |
400 lemma rfv_ty_rsp[quot_respect]: |
401 lemma rfv_rty_rsp[quot_respect]: |
401 "(aty ===> op =) rfv_ty rfv_ty" |
402 "(arty ===> op =) rfv_rty rfv_rty" |
402 by (simp add: alpha_rfv) |
403 by (simp add: alpha_rfv) |
403 lemma rfv_kind_rsp[quot_respect]: |
404 lemma rfv_rkind_rsp[quot_respect]: |
404 "(akind ===> op =) rfv_kind rfv_kind" |
405 "(arkind ===> op =) rfv_rkind rfv_rkind" |
405 by (simp add: alpha_rfv) |
406 by (simp add: alpha_rfv) |
406 lemma rfv_trm_rsp[quot_respect]: |
407 lemma rfv_rtrm_rsp[quot_respect]: |
407 "(atrm ===> op =) rfv_trm rfv_trm" |
408 "(artrm ===> op =) rfv_rtrm rfv_rtrm" |
408 by (simp add: alpha_rfv) |
409 by (simp add: alpha_rfv) |
409 |
410 |
410 thm kind_ty_trm.induct |
411 thm rkind_rty_rtrm.induct |
411 lemmas KIND_TY_TRM_induct = kind_ty_trm.induct[quot_lifted] |
412 lemmas RKIND_RTY_RTRM_induct = rkind_rty_rtrm.induct[quot_lifted] |
412 |
413 |
413 thm kind_ty_trm.inducts |
414 thm rkind_rty_rtrm.inducts |
414 lemmas KIND_TY_TRM_inducts = kind_ty_trm.inducts[quot_lifted] |
415 lemmas RKIND_RTY_RTRM_inducts = rkind_rty_rtrm.inducts[quot_lifted] |
415 |
416 |
416 instantiation KIND and TY and TRM :: pt |
417 instantiation RKIND and RTY and RTRM :: pt |
417 begin |
418 begin |
418 |
419 |
419 quotient_definition |
420 quotient_definition |
420 "permute_KIND :: perm \<Rightarrow> KIND \<Rightarrow> KIND" |
421 "permute_RKIND :: perm \<Rightarrow> RKIND \<Rightarrow> RKIND" |
421 is |
422 is |
422 "permute :: perm \<Rightarrow> kind \<Rightarrow> kind" |
423 "permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind" |
423 |
424 |
424 quotient_definition |
425 quotient_definition |
425 "permute_TY :: perm \<Rightarrow> TY \<Rightarrow> TY" |
426 "permute_RTY :: perm \<Rightarrow> RTY \<Rightarrow> RTY" |
426 is |
427 is |
427 "permute :: perm \<Rightarrow> ty \<Rightarrow> ty" |
428 "permute :: perm \<Rightarrow> rty \<Rightarrow> rty" |
428 |
429 |
429 quotient_definition |
430 quotient_definition |
430 "permute_TRM :: perm \<Rightarrow> TRM \<Rightarrow> TRM" |
431 "permute_RTRM :: perm \<Rightarrow> RTRM \<Rightarrow> RTRM" |
431 is |
432 is |
432 "permute :: perm \<Rightarrow> trm \<Rightarrow> trm" |
433 "permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm" |
433 |
434 |
434 lemmas permute_ktt[simp] = permute_kind_permute_ty_permute_trm.simps[quot_lifted] |
435 lemmas permute_ktt[simp] = permute_rkind_permute_rty_permute_rtrm.simps[quot_lifted] |
435 |
436 |
436 lemma perm_zero_ok: "0 \<bullet> (x :: KIND) = x \<and> 0 \<bullet> (y :: TY) = y \<and> 0 \<bullet> (z :: TRM) = z" |
437 lemma perm_zero_ok: "0 \<bullet> (x :: RKIND) = x \<and> 0 \<bullet> (y :: RTY) = y \<and> 0 \<bullet> (z :: RTRM) = z" |
437 apply (induct rule: KIND_TY_TRM_induct) |
438 apply (induct rule: RKIND_RTY_RTRM_induct) |
438 apply (simp_all) |
439 apply (simp_all) |
439 done |
440 done |
440 |
441 |
441 lemma perm_add_ok: |
442 lemma perm_add_ok: |
442 "((p + q) \<bullet> (x1 :: KIND) = (p \<bullet> q \<bullet> x1))" |
443 "((p + q) \<bullet> (x1 :: RKIND) = (p \<bullet> q \<bullet> x1))" |
443 "((p + q) \<bullet> (x2 :: TY) = p \<bullet> q \<bullet> x2)" |
444 "((p + q) \<bullet> (x2 :: RTY) = p \<bullet> q \<bullet> x2)" |
444 "((p + q) \<bullet> (x3 :: TRM) = p \<bullet> q \<bullet> x3)" |
445 "((p + q) \<bullet> (x3 :: RTRM) = p \<bullet> q \<bullet> x3)" |
445 apply (induct x1 and x2 and x3 rule: KIND_TY_TRM_inducts) |
446 apply (induct x1 and x2 and x3 rule: RKIND_RTY_RTRM_inducts) |
446 apply (simp_all) |
447 apply (simp_all) |
447 done |
448 done |
448 |
449 |
449 instance |
450 instance |
450 apply default |
451 apply default |
451 apply (simp_all add: perm_zero_ok perm_add_ok) |
452 apply (simp_all add: perm_zero_ok perm_add_ok) |
452 done |
453 done |
453 |
454 |
454 end |
455 end |
455 |
456 |
456 lemmas AKIND_ATY_ATRM_inducts = akind_aty_atrm.inducts[unfolded alpha_gen, quot_lifted, folded alpha_gen] |
457 lemmas ARKIND_ARTY_ARTRM_inducts = arkind_arty_artrm.inducts[unfolded alpha_gen, quot_lifted, folded alpha_gen] |
457 |
458 |
458 lemmas KIND_TY_TRM_INJECT = akind_aty_atrm_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] |
459 lemmas RKIND_RTY_RTRM_INJECT = arkind_arty_artrm_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] |
459 |
460 |
460 lemmas fv_kind_ty_trm = rfv_kind_rfv_ty_rfv_trm.simps[quot_lifted] |
461 lemmas fv_rkind_rty_rtrm = rfv_rkind_rfv_rty_rfv_rtrm.simps[quot_lifted] |
461 |
462 |
462 lemmas fv_eqvt = rfv_eqvt[quot_lifted] |
463 lemmas fv_eqvt = rfv_eqvt[quot_lifted] |
463 |
464 |
464 lemma supp_kind_ty_trm_easy: |
465 lemma supp_rkind_rty_rtrm_easy: |
465 "supp TYP = {}" |
466 "supp TYP = {}" |
466 "supp (TCONST i) = {atom i}" |
467 "supp (TCONST i) = {atom i}" |
467 "supp (TAPP A M) = supp A \<union> supp M" |
468 "supp (TAPP A M) = supp A \<union> supp M" |
468 "supp (CONS i) = {atom i}" |
469 "supp (CONS i) = {atom i}" |
469 "supp (VAR x) = {atom x}" |
470 "supp (VAR x) = {atom x}" |
470 "supp (APP M N) = supp M \<union> supp N" |
471 "supp (APP M N) = supp M \<union> supp N" |
471 apply (simp_all add: supp_def permute_ktt KIND_TY_TRM_INJECT) |
472 apply (simp_all add: supp_def permute_ktt RKIND_RTY_RTRM_INJECT) |
472 apply (simp_all only: supp_at_base[simplified supp_def]) |
473 apply (simp_all only: supp_at_base[simplified supp_def]) |
473 apply (simp_all add: Collect_imp_eq Collect_neg_eq) |
474 apply (simp_all add: Collect_imp_eq Collect_neg_eq) |
474 done |
475 done |
475 |
476 |
476 lemma supp_bind: |
477 lemma supp_bind: |
477 "(supp (atom na, (ty, ki))) supports (KPI ty na ki)" |
478 "(supp (atom na, (ty, ki))) supports (KPI ty na ki)" |
478 "(supp (atom na, (ty, ty2))) supports (TPI ty na ty2)" |
479 "(supp (atom na, (ty, ty2))) supports (TPI ty na ty2)" |
479 "(supp (atom na, (ty, trm))) supports (LAM ty na trm)" |
480 "(supp (atom na, (ty, rtrm))) supports (LAM ty na rtrm)" |
480 apply(simp_all add: supports_def) |
481 apply(simp_all add: supports_def) |
481 apply(fold fresh_def) |
482 apply(fold fresh_def) |
482 apply(simp_all add: fresh_Pair swap_fresh_fresh) |
483 apply(simp_all add: fresh_Pair swap_fresh_fresh) |
483 apply(clarify) |
484 apply(clarify) |
484 apply(subst swap_at_base_simps(3)) |
485 apply(subst swap_at_base_simps(3)) |
506 apply(rule supports_finite) |
507 apply(rule supports_finite) |
507 apply(rule supp_bind(3)) |
508 apply(rule supp_bind(3)) |
508 apply(simp add: supp_Pair supp_atom) |
509 apply(simp add: supp_Pair supp_atom) |
509 done |
510 done |
510 |
511 |
511 instance KIND and TY and TRM :: fs |
512 instance RKIND and RTY and RTRM :: fs |
512 apply(default) |
513 apply(default) |
513 apply(simp_all only: KIND_TY_TRM_fs) |
514 apply(simp_all only: RKIND_RTY_RTRM_fs) |
514 done |
515 done |
515 |
516 |
516 lemma supp_fv: |
517 lemma supp_fv: |
517 "supp t1 = fv_kind t1" |
518 "supp t1 = fv_rkind t1" |
518 "supp t2 = fv_ty t2" |
519 "supp t2 = fv_rty t2" |
519 "supp t3 = fv_trm t3" |
520 "supp t3 = fv_rtrm t3" |
520 apply(induct t1 and t2 and t3 rule: KIND_TY_TRM_inducts) |
521 apply(induct t1 and t2 and t3 rule: RKIND_RTY_RTRM_inducts) |
521 apply (simp_all add: supp_kind_ty_trm_easy) |
522 apply (simp_all add: supp_rkind_rty_rtrm_easy) |
522 apply (simp_all add: fv_kind_ty_trm) |
523 apply (simp_all add: fv_rkind_rty_rtrm) |
523 apply(subgoal_tac "supp (KPI ty name kind) = supp ty \<union> supp (Abs {atom name} kind)") |
524 apply(subgoal_tac "supp (KPI rty name rkind) = supp rty \<union> supp (Abs {atom name} rkind)") |
524 apply(simp add: supp_Abs Set.Un_commute) |
525 apply(simp add: supp_Abs Set.Un_commute) |
525 apply(simp (no_asm) add: supp_def) |
526 apply(simp (no_asm) add: supp_def) |
526 apply(simp add: KIND_TY_TRM_INJECT) |
527 apply(simp add: RKIND_RTY_RTRM_INJECT) |
527 apply(simp add: Abs_eq_iff) |
528 apply(simp add: Abs_eq_iff) |
528 apply(simp add: alpha_gen) |
529 apply(simp add: alpha_gen) |
529 apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute insert_eqvt empty_eqvt atom_eqvt) |
530 apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute insert_eqvt empty_eqvt atom_eqvt) |
530 apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric]) |
531 apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric]) |
531 apply(subgoal_tac "supp (TPI ty1 name ty2) = supp ty1 \<union> supp (Abs {atom name} ty2)") |
532 apply(subgoal_tac "supp (TPI rty1 name rty2) = supp rty1 \<union> supp (Abs {atom name} rty2)") |
532 apply(simp add: supp_Abs Set.Un_commute) |
533 apply(simp add: supp_Abs Set.Un_commute) |
533 apply(simp (no_asm) add: supp_def) |
534 apply(simp (no_asm) add: supp_def) |
534 apply(simp add: KIND_TY_TRM_INJECT) |
535 apply(simp add: RKIND_RTY_RTRM_INJECT) |
535 apply(simp add: Abs_eq_iff) |
536 apply(simp add: Abs_eq_iff) |
536 apply(simp add: alpha_gen) |
537 apply(simp add: alpha_gen) |
537 apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric] insert_eqvt empty_eqvt atom_eqvt) |
538 apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric] insert_eqvt empty_eqvt atom_eqvt) |
538 apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute) |
539 apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute) |
539 apply(subgoal_tac "supp (LAM ty name trm) = supp ty \<union> supp (Abs {atom name} trm)") |
540 apply(subgoal_tac "supp (LAM rty name rtrm) = supp rty \<union> supp (Abs {atom name} rtrm)") |
540 apply(simp add: supp_Abs Set.Un_commute) |
541 apply(simp add: supp_Abs Set.Un_commute) |
541 apply(simp (no_asm) add: supp_def) |
542 apply(simp (no_asm) add: supp_def) |
542 apply(simp add: KIND_TY_TRM_INJECT) |
543 apply(simp add: RKIND_RTY_RTRM_INJECT) |
543 apply(simp add: Abs_eq_iff) |
544 apply(simp add: Abs_eq_iff) |
544 apply(simp add: alpha_gen) |
545 apply(simp add: alpha_gen) |
545 apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric] insert_eqvt empty_eqvt atom_eqvt) |
546 apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric] insert_eqvt empty_eqvt atom_eqvt) |
546 apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute) |
547 apply(simp add: Collect_imp_eq Collect_neg_eq Set.Un_commute) |
547 done |
548 done |
548 |
549 |
549 (* Not needed anymore *) |
550 (* Not needed anymore *) |
550 lemma supp_kpi_pre: "supp (KPI A x K) = (supp (Abs {atom x} K)) \<union> supp A" |
551 lemma supp_kpi_pre: "supp (KPI A x K) = (supp (Abs {atom x} K)) \<union> supp A" |
551 apply (simp add: permute_set_eq supp_def Abs_eq_iff KIND_TY_TRM_INJECT) |
552 apply (simp add: permute_set_eq supp_def Abs_eq_iff RKIND_RTY_RTRM_INJECT) |
552 apply (simp add: alpha_gen supp_fv) |
553 apply (simp add: alpha_gen supp_fv) |
553 apply (simp add: Collect_imp_eq Collect_neg_eq add: atom_eqvt Set.Un_commute) |
554 apply (simp add: Collect_imp_eq Collect_neg_eq add: atom_eqvt Set.Un_commute) |
554 done |
555 done |
555 |
556 |
556 lemma supp_kind_ty_trm: |
557 lemma supp_rkind_rty_rtrm: |
557 "supp TYP = {}" |
558 "supp TYP = {}" |
558 "supp (KPI A x K) = supp A \<union> (supp K - {atom x})" |
559 "supp (KPI A x K) = supp A \<union> (supp K - {atom x})" |
559 "supp (TCONST i) = {atom i}" |
560 "supp (TCONST i) = {atom i}" |
560 "supp (TAPP A M) = supp A \<union> supp M" |
561 "supp (TAPP A M) = supp A \<union> supp M" |
561 "supp (TPI A x B) = supp A \<union> (supp B - {atom x})" |
562 "supp (TPI A x B) = supp A \<union> (supp B - {atom x})" |
562 "supp (CONS i) = {atom i}" |
563 "supp (CONS i) = {atom i}" |
563 "supp (VAR x) = {atom x}" |
564 "supp (VAR x) = {atom x}" |
564 "supp (APP M N) = supp M \<union> supp N" |
565 "supp (APP M N) = supp M \<union> supp N" |
565 "supp (LAM A x M) = supp A \<union> (supp M - {atom x})" |
566 "supp (LAM A x M) = supp A \<union> (supp M - {atom x})" |
566 apply (simp_all only: supp_kind_ty_trm_easy) |
567 apply (simp_all only: supp_rkind_rty_rtrm_easy) |
567 apply (simp_all only: supp_fv fv_kind_ty_trm) |
568 apply (simp_all only: supp_fv fv_rkind_rty_rtrm) |
568 done |
569 done |
569 |
570 |
570 end |
571 end |
571 |
572 |
572 |
573 |