188 assumes q1: "Quotient R1 abs1 rep1" |
189 assumes q1: "Quotient R1 abs1 rep1" |
189 and q2: "Quotient R2 abs2 rep2" |
190 and q2: "Quotient R2 abs2 rep2" |
190 shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)" |
191 shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)" |
191 proof - |
192 proof - |
192 have "\<forall>a. (rep1 ---> abs2) ((abs1 ---> rep2) a) = a" |
193 have "\<forall>a. (rep1 ---> abs2) ((abs1 ---> rep2) a) = a" |
193 apply(simp add: expand_fun_eq) |
194 using q1 q2 |
194 using q1 q2 |
195 unfolding Quotient_def |
195 apply(simp add: Quotient_def) |
196 unfolding expand_fun_eq |
196 done |
197 by simp |
197 moreover |
198 moreover |
198 have "\<forall>a. (R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)" |
199 have "\<forall>a. (R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)" |
199 apply(auto) |
200 using q1 q2 |
200 using q1 q2 unfolding Quotient_def |
201 unfolding Quotient_def |
201 apply(metis) |
202 by (simp (no_asm)) (metis) |
202 done |
|
203 moreover |
203 moreover |
204 have "\<forall>r s. (R1 ===> R2) r s = ((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and> |
204 have "\<forall>r s. (R1 ===> R2) r s = ((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and> |
205 (rep1 ---> abs2) r = (rep1 ---> abs2) s)" |
205 (rep1 ---> abs2) r = (rep1 ---> abs2) s)" |
206 apply(auto simp add: expand_fun_eq) |
206 unfolding expand_fun_eq |
|
207 apply(auto) |
207 using q1 q2 unfolding Quotient_def |
208 using q1 q2 unfolding Quotient_def |
208 apply(metis) |
209 apply(metis) |
209 using q1 q2 unfolding Quotient_def |
210 using q1 q2 unfolding Quotient_def |
210 apply(metis) |
211 apply(metis) |
211 using q1 q2 unfolding Quotient_def |
212 using q1 q2 unfolding Quotient_def |
219 qed |
220 qed |
220 |
221 |
221 lemma abs_o_rep: |
222 lemma abs_o_rep: |
222 assumes a: "Quotient R Abs Rep" |
223 assumes a: "Quotient R Abs Rep" |
223 shows "Abs o Rep = id" |
224 shows "Abs o Rep = id" |
224 apply(rule ext) |
225 unfolding expand_fun_eq |
225 apply(simp add: Quotient_abs_rep[OF a]) |
226 by (simp add: Quotient_abs_rep[OF a]) |
226 done |
|
227 |
227 |
228 lemma equals_rsp: |
228 lemma equals_rsp: |
229 assumes q: "Quotient R Abs Rep" |
229 assumes q: "Quotient R Abs Rep" |
230 and a: "R xa xb" "R ya yb" |
230 and a: "R xa xb" "R ya yb" |
231 shows "R xa ya = R xb yb" |
231 shows "R xa ya = R xb yb" |
232 using Quotient_symp[OF q] Quotient_transp[OF q] unfolding symp_def transp_def |
232 using a Quotient_symp[OF q] Quotient_transp[OF q] |
233 using a by blast |
233 unfolding symp_def transp_def |
|
234 by blast |
234 |
235 |
235 lemma lambda_prs: |
236 lemma lambda_prs: |
236 assumes q1: "Quotient R1 Abs1 Rep1" |
237 assumes q1: "Quotient R1 Abs1 Rep1" |
237 and q2: "Quotient R2 Abs2 Rep2" |
238 and q2: "Quotient R2 Abs2 Rep2" |
238 shows "(Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) = (\<lambda>x. f x)" |
239 shows "(Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) = (\<lambda>x. f x)" |
250 |
251 |
251 lemma rep_abs_rsp: |
252 lemma rep_abs_rsp: |
252 assumes q: "Quotient R Abs Rep" |
253 assumes q: "Quotient R Abs Rep" |
253 and a: "R x1 x2" |
254 and a: "R x1 x2" |
254 shows "R x1 (Rep (Abs x2))" |
255 shows "R x1 (Rep (Abs x2))" |
255 using a |
256 using a Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q] |
256 by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]) |
257 by metis |
257 |
258 |
258 lemma rep_abs_rsp_left: |
259 lemma rep_abs_rsp_left: |
259 assumes q: "Quotient R Abs Rep" |
260 assumes q: "Quotient R Abs Rep" |
260 and a: "R x1 x2" |
261 and a: "R x1 x2" |
261 shows "R (Rep (Abs x1)) x2" |
262 shows "R (Rep (Abs x1)) x2" |
262 using a |
263 using a Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q] |
263 by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]) |
264 by metis |
264 |
265 |
265 (* In the following theorem R1 can be instantiated with anything, |
266 text{* |
266 but we know some of the types of the Rep and Abs functions; |
267 In the following theorem R1 can be instantiated with anything, |
267 so by solving Quotient assumptions we can get a unique R1 that |
268 but we know some of the types of the Rep and Abs functions; |
268 will be provable; which is why we need to use apply_rsp and |
269 so by solving Quotient assumptions we can get a unique R1 that |
269 not the primed version *) |
270 will be provable; which is why we need to use apply_rsp and |
|
271 not the primed version *} |
|
272 |
270 lemma apply_rsp: |
273 lemma apply_rsp: |
271 fixes f g::"'a \<Rightarrow> 'c" |
274 fixes f g::"'a \<Rightarrow> 'c" |
272 assumes q: "Quotient R1 Abs1 Rep1" |
275 assumes q: "Quotient R1 Abs1 Rep1" |
273 and a: "(R1 ===> R2) f g" "R1 x y" |
276 and a: "(R1 ===> R2) f g" "R1 x y" |
274 shows "R2 (f x) (g y)" |
277 shows "R2 (f x) (g y)" |
283 |
286 |
284 lemma ball_reg_eqv: |
287 lemma ball_reg_eqv: |
285 fixes P :: "'a \<Rightarrow> bool" |
288 fixes P :: "'a \<Rightarrow> bool" |
286 assumes a: "equivp R" |
289 assumes a: "equivp R" |
287 shows "Ball (Respects R) P = (All P)" |
290 shows "Ball (Respects R) P = (All P)" |
288 by (metis equivp_def in_respects a) |
291 using a |
|
292 unfolding equivp_def |
|
293 by (auto simp add: in_respects) |
289 |
294 |
290 lemma bex_reg_eqv: |
295 lemma bex_reg_eqv: |
291 fixes P :: "'a \<Rightarrow> bool" |
296 fixes P :: "'a \<Rightarrow> bool" |
292 assumes a: "equivp R" |
297 assumes a: "equivp R" |
293 shows "Bex (Respects R) P = (Ex P)" |
298 shows "Bex (Respects R) P = (Ex P)" |
294 by (metis equivp_def in_respects a) |
299 using a |
|
300 unfolding equivp_def |
|
301 by (auto simp add: in_respects) |
295 |
302 |
296 lemma ball_reg_right: |
303 lemma ball_reg_right: |
297 assumes a: "\<And>x. R x \<Longrightarrow> P x \<longrightarrow> Q x" |
304 assumes a: "\<And>x. R x \<Longrightarrow> P x \<longrightarrow> Q x" |
298 shows "All P \<longrightarrow> Ball R Q" |
305 shows "All P \<longrightarrow> Ball R Q" |
299 by (metis COMBC_def Collect_def Collect_mem_eq a) |
306 using a by (metis COMBC_def Collect_def Collect_mem_eq) |
300 |
307 |
301 lemma bex_reg_left: |
308 lemma bex_reg_left: |
302 assumes a: "\<And>x. R x \<Longrightarrow> Q x \<longrightarrow> P x" |
309 assumes a: "\<And>x. R x \<Longrightarrow> Q x \<longrightarrow> P x" |
303 shows "Bex R Q \<longrightarrow> Ex P" |
310 shows "Bex R Q \<longrightarrow> Ex P" |
304 by (metis COMBC_def Collect_def Collect_mem_eq a) |
311 using a by (metis COMBC_def Collect_def Collect_mem_eq) |
305 |
312 |
306 lemma ball_reg_left: |
313 lemma ball_reg_left: |
307 assumes a: "equivp R" |
314 assumes a: "equivp R" |
308 shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ball (Respects R) Q \<longrightarrow> All P" |
315 shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ball (Respects R) Q \<longrightarrow> All P" |
309 by (metis equivp_reflp in_respects a) |
316 using a by (metis equivp_reflp in_respects) |
310 |
317 |
311 lemma bex_reg_right: |
318 lemma bex_reg_right: |
312 assumes a: "equivp R" |
319 assumes a: "equivp R" |
313 shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ex Q \<longrightarrow> Bex (Respects R) P" |
320 shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ex Q \<longrightarrow> Bex (Respects R) P" |
314 by (metis equivp_reflp in_respects a) |
321 using a by (metis equivp_reflp in_respects) |
315 |
322 |
316 lemma ball_reg_eqv_range: |
323 lemma ball_reg_eqv_range: |
317 fixes P::"'a \<Rightarrow> bool" |
324 fixes P::"'a \<Rightarrow> bool" |
318 and x::"'a" |
325 and x::"'a" |
319 assumes a: "equivp R2" |
326 assumes a: "equivp R2" |
364 and b: "Bex R P" |
371 and b: "Bex R P" |
365 shows "Bex R Q" |
372 shows "Bex R Q" |
366 using a b by (metis COMBC_def Collect_def Collect_mem_eq) |
373 using a b by (metis COMBC_def Collect_def Collect_mem_eq) |
367 |
374 |
368 lemma ball_all_comm: |
375 lemma ball_all_comm: |
369 "(\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)) \<Longrightarrow> ((\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y))" |
376 assumes "\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)" |
370 by auto |
377 shows "(\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y)" |
|
378 using assms by auto |
371 |
379 |
372 lemma bex_ex_comm: |
380 lemma bex_ex_comm: |
373 "((\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)) \<Longrightarrow> ((\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y))" |
381 assumes "(\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)" |
374 by auto |
382 shows "(\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y)" |
|
383 using assms by auto |
375 |
384 |
376 section {* Bounded abstraction *} |
385 section {* Bounded abstraction *} |
377 |
386 |
378 definition |
387 definition |
379 Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" |
388 Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" |
382 |
391 |
383 lemma babs_rsp: |
392 lemma babs_rsp: |
384 assumes q: "Quotient R1 Abs1 Rep1" |
393 assumes q: "Quotient R1 Abs1 Rep1" |
385 and a: "(R1 ===> R2) f g" |
394 and a: "(R1 ===> R2) f g" |
386 shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)" |
395 shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)" |
387 apply (auto simp add: Babs_def) |
396 apply (auto simp add: Babs_def in_respects) |
388 apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1") |
397 apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1") |
389 using a apply (simp add: Babs_def) |
398 using a apply (simp add: Babs_def) |
390 apply (simp add: in_respects) |
399 apply (simp add: in_respects) |
391 using Quotient_rel[OF q] |
400 using Quotient_rel[OF q] |
392 by metis |
401 by metis |