387 by (metis COMBC_def Collect_def Collect_mem_eq in_respects fun_map.simps id_apply) |
399 by (metis COMBC_def Collect_def Collect_mem_eq in_respects fun_map.simps id_apply) |
388 |
400 |
389 lemma fun_rel_id: |
401 lemma fun_rel_id: |
390 assumes a: "\<And>x y. R1 x y \<Longrightarrow> R2 (f x) (g y)" |
402 assumes a: "\<And>x y. R1 x y \<Longrightarrow> R2 (f x) (g y)" |
391 shows "(R1 ===> R2) f g" |
403 shows "(R1 ===> R2) f g" |
392 using a by simp |
404 using a by simp |
393 |
405 |
394 lemma fun_rel_id_asm: |
406 lemma fun_rel_id_asm: |
395 assumes a: "\<And>x y. R1 x y \<Longrightarrow> (A \<longrightarrow> R2 (f x) (g y))" |
407 assumes a: "\<And>x y. R1 x y \<Longrightarrow> (A \<longrightarrow> R2 (f x) (g y))" |
396 shows "A \<longrightarrow> (R1 ===> R2) f g" |
408 shows "A \<longrightarrow> (R1 ===> R2) f g" |
397 using a by auto |
409 using a by auto |
398 |
410 |
399 lemma quot_rel_rsp: |
411 lemma quot_rel_rsp: |
400 assumes a: "Quotient R Abs Rep" |
412 assumes a: "Quotient R Abs Rep" |
401 shows "(R ===> R ===> op =) R R" |
413 shows "(R ===> R ===> op =) R R" |
402 apply(rule fun_rel_id)+ |
414 apply(rule fun_rel_id)+ |
403 apply(rule equals_rsp[OF a]) |
415 apply(rule equals_rsp[OF a]) |
404 apply(assumption)+ |
416 apply(assumption)+ |
405 done |
417 done |
406 |
418 |
407 |
419 lemma o_prs: |
|
420 assumes q1: "Quotient R1 Abs1 Rep1" |
|
421 and q2: "Quotient R2 Abs2 Rep2" |
|
422 and q3: "Quotient R3 Abs3 Rep3" |
|
423 shows "(Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g)) = f o g" |
|
424 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3] |
|
425 unfolding o_def expand_fun_eq by simp |
|
426 |
|
427 lemma o_rsp: |
|
428 assumes q1: "Quotient R1 Abs1 Rep1" |
|
429 and q2: "Quotient R2 Abs2 Rep2" |
|
430 and q3: "Quotient R3 Abs3 Rep3" |
|
431 and a1: "(R2 ===> R3) f1 f2" |
|
432 and a2: "(R1 ===> R2) g1 g2" |
|
433 shows "(R1 ===> R3) (f1 o g1) (f2 o g2)" |
|
434 using a1 a2 unfolding o_def expand_fun_eq |
|
435 by (auto) |
|
436 |
|
437 lemma cond_prs: |
|
438 assumes a: "Quotient R absf repf" |
|
439 shows "absf (if a then repf b else repf c) = (if a then b else c)" |
|
440 using a unfolding Quotient_def by auto |
|
441 |
|
442 lemma if_prs: |
|
443 assumes q: "Quotient R Abs Rep" |
|
444 shows "Abs (If a (Rep b) (Rep c)) = If a b c" |
|
445 using Quotient_abs_rep[OF q] by auto |
|
446 |
|
447 (* q not used *) |
|
448 lemma if_rsp: |
|
449 assumes q: "Quotient R Abs Rep" |
|
450 and a: "a1 = a2" "R b1 b2" "R c1 c2" |
|
451 shows "R (If a1 b1 c1) (If a2 b2 c2)" |
|
452 using a by auto |
|
453 |
|
454 lemma let_prs: |
|
455 assumes q1: "Quotient R1 Abs1 Rep1" |
|
456 and q2: "Quotient R2 Abs2 Rep2" |
|
457 shows "Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f)) = Let x f" |
|
458 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto |
|
459 |
|
460 lemma let_rsp: |
|
461 assumes q1: "Quotient R1 Abs1 Rep1" |
|
462 and a1: "(R1 ===> R2) f g" |
|
463 and a2: "R1 x y" |
|
464 shows "R2 ((Let x f)::'c) ((Let y g)::'c)" |
|
465 using apply_rsp[OF q1 a1] a2 by auto |
408 |
466 |
409 |
467 |
410 |
468 |
411 |
469 |
412 (******************************************) |
470 (******************************************) |
413 (* REST OF THE FILE IS UNUSED (until now) *) |
471 (* REST OF THE FILE IS UNUSED (until now) *) |
414 (******************************************) |
472 (******************************************) |
415 |
473 |
416 (* Always safe to apply, but not too helpful *) |
|
417 lemma app_prs2: |
|
418 assumes q1: "Quotient R1 abs1 rep1" |
|
419 and q2: "Quotient R2 abs2 rep2" |
|
420 shows "((abs1 ---> rep2) ((rep1 ---> abs2) f) (rep1 x)) = rep2 (((rep1 ---> abs2) f) x)" |
|
421 unfolding expand_fun_eq |
|
422 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] |
|
423 by simp |
|
424 |
|
425 lemma Quotient_rel_abs: |
|
426 assumes a: "Quotient E Abs Rep" |
|
427 shows "E r s \<Longrightarrow> Abs r = Abs s" |
|
428 using a unfolding Quotient_def |
|
429 by blast |
|
430 |
|
431 lemma Quotient_rel_abs_eq: |
|
432 assumes a: "Quotient E Abs Rep" |
|
433 shows "E r r \<Longrightarrow> E s s \<Longrightarrow> E r s = (Abs r = Abs s)" |
|
434 using a unfolding Quotient_def |
|
435 by blast |
|
436 |
|
437 lemma in_fun: |
474 lemma in_fun: |
438 shows "x \<in> ((f ---> g) s) = g (f x \<in> s)" |
475 shows "x \<in> ((f ---> g) s) = g (f x \<in> s)" |
439 by (simp add: mem_def) |
476 by (simp add: mem_def) |
440 |
477 |
441 lemma RESPECTS_THM: |
478 lemma respects_thm: |
442 shows "Respects (R1 ===> R2) f = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (f y))" |
479 shows "Respects (R1 ===> R2) f = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (f y))" |
443 unfolding Respects_def |
480 unfolding Respects_def |
444 by (simp add: expand_fun_eq) |
481 by (simp add: expand_fun_eq) |
445 |
482 |
446 lemma RESPECTS_REP_ABS: |
483 lemma respects_rep_abs: |
447 assumes a: "Quotient R1 Abs1 Rep1" |
484 assumes a: "Quotient R1 Abs1 Rep1" |
448 and b: "Respects (R1 ===> R2) f" |
485 and b: "Respects (R1 ===> R2) f" |
449 and c: "R1 x x" |
486 and c: "R1 x x" |
450 shows "R2 (f (Rep1 (Abs1 x))) (f x)" |
487 shows "R2 (f (Rep1 (Abs1 x))) (f x)" |
451 using a b[simplified RESPECTS_THM] c unfolding Quotient_def |
488 using a b[simplified respects_thm] c unfolding Quotient_def |
452 by blast |
489 by blast |
453 |
490 |
454 lemma RESPECTS_MP: |
491 lemma respects_mp: |
455 assumes a: "Respects (R1 ===> R2) f" |
492 assumes a: "Respects (R1 ===> R2) f" |
456 and b: "R1 x y" |
493 and b: "R1 x y" |
457 shows "R2 (f x) (f y)" |
494 shows "R2 (f x) (f y)" |
458 using a b unfolding Respects_def |
495 using a b unfolding Respects_def |
459 by simp |
496 by simp |
460 |
497 |
461 lemma RESPECTS_o: |
498 lemma respects_o: |
462 assumes a: "Respects (R2 ===> R3) f" |
499 assumes a: "Respects (R2 ===> R3) f" |
463 and b: "Respects (R1 ===> R2) g" |
500 and b: "Respects (R1 ===> R2) g" |
464 shows "Respects (R1 ===> R3) (f o g)" |
501 shows "Respects (R1 ===> R3) (f o g)" |
465 using a b unfolding Respects_def |
502 using a b unfolding Respects_def |
466 by simp |
503 by simp |
467 |
504 |
468 lemma fun_rel_EQ_REL: |
505 lemma fun_rel_eq_rel: |
469 assumes q1: "Quotient R1 Abs1 Rep1" |
506 assumes q1: "Quotient R1 Abs1 Rep1" |
470 and q2: "Quotient R2 Abs2 Rep2" |
507 and q2: "Quotient R2 Abs2 Rep2" |
471 shows "(R1 ===> R2) f g = ((Respects (R1 ===> R2) f) \<and> (Respects (R1 ===> R2) g) |
508 shows "(R1 ===> R2) f g = ((Respects (R1 ===> R2) f) \<and> (Respects (R1 ===> R2) g) |
472 \<and> ((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g))" |
509 \<and> ((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g))" |
473 using fun_quotient[OF q1 q2] unfolding Respects_def Quotient_def expand_fun_eq |
510 using fun_quotient[OF q1 q2] unfolding Respects_def Quotient_def expand_fun_eq |
474 by blast |
511 by blast |
475 |
512 |
476 (* Not used since in the end we just unfold fun_map *) |
513 lemma let_babs: |
477 lemma APP_PRS: |
514 "v \<in> r \<Longrightarrow> Let v (Babs r lam) = Let v lam" |
478 assumes q1: "Quotient R1 abs1 rep1" |
515 by (simp add: Babs_def) |
479 and q2: "Quotient R2 abs2 rep2" |
516 |
480 shows "abs2 ((abs1 ---> rep2) f (rep1 x)) = f x" |
517 lemma fun_rel_equals: |
481 unfolding expand_fun_eq |
|
482 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] |
|
483 by simp |
|
484 |
|
485 (* Ask Peter: assumption q1 and q2 not used and lemma is the 'identity' *) |
|
486 lemma LAMBDA_RSP: |
|
487 assumes q1: "Quotient R1 Abs1 Rep1" |
|
488 and q2: "Quotient R2 Abs2 Rep2" |
|
489 and a: "(R1 ===> R2) f1 f2" |
|
490 shows "(R1 ===> R2) (\<lambda>x. f1 x) (\<lambda>y. f2 y)" |
|
491 by (rule a) |
|
492 |
|
493 (* ASK Peter about next four lemmas in quotientScript |
|
494 lemma ABSTRACT_PRS: |
|
495 assumes q1: "Quotient R1 Abs1 Rep1" |
|
496 and q2: "Quotient R2 Abs2 Rep2" |
|
497 shows "f = (Rep1 ---> Abs2) ???" |
|
498 *) |
|
499 |
|
500 |
|
501 lemma fun_rel_EQUALS: |
|
502 assumes q1: "Quotient R1 Abs1 Rep1" |
518 assumes q1: "Quotient R1 Abs1 Rep1" |
503 and q2: "Quotient R2 Abs2 Rep2" |
519 and q2: "Quotient R2 Abs2 Rep2" |
504 and r1: "Respects (R1 ===> R2) f" |
520 and r1: "Respects (R1 ===> R2) f" |
505 and r2: "Respects (R1 ===> R2) g" |
521 and r2: "Respects (R1 ===> R2) g" |
506 shows "((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g) = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (g y))" |
522 shows "((Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g) = (\<forall>x y. R1 x y \<longrightarrow> R2 (f x) (g y))" |
507 apply(rule_tac iffI) |
523 apply(rule_tac iffI) |
508 using fun_quotient[OF q1 q2] r1 r2 unfolding Quotient_def Respects_def |
524 using fun_quotient[OF q1 q2] r1 r2 unfolding Quotient_def Respects_def |
509 apply(metis apply_rsp') |
525 apply(metis apply_rsp') |
510 using r1 unfolding Respects_def expand_fun_eq |
526 using r1 unfolding Respects_def expand_fun_eq |
511 apply(simp (no_asm_use)) |
527 apply(simp (no_asm_use)) |
512 apply(metis Quotient_rel[OF q2] Quotient_rel_rep[OF q1]) |
528 apply(metis Quotient_rel[OF q2] Quotient_rel_rep[OF q1]) |
513 done |
529 done |
514 |
530 |
515 (* ask Peter: fun_rel_IMP used twice *) |
531 (* ask Peter: fun_rel_IMP used twice *) |
516 lemma fun_rel_IMP2: |
532 lemma fun_rel_IMP2: |
517 assumes q1: "Quotient R1 Abs1 Rep1" |
533 assumes q1: "Quotient R1 Abs1 Rep1" |
518 and q2: "Quotient R2 Abs2 Rep2" |
534 and q2: "Quotient R2 Abs2 Rep2" |
519 and r1: "Respects (R1 ===> R2) f" |
535 and r1: "Respects (R1 ===> R2) f" |
520 and r2: "Respects (R1 ===> R2) g" |
536 and r2: "Respects (R1 ===> R2) g" |
521 and a: "(Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g" |
537 and a: "(Rep1 ---> Abs2) f = (Rep1 ---> Abs2) g" |
522 shows "R1 x y \<Longrightarrow> R2 (f x) (g y)" |
538 shows "R1 x y \<Longrightarrow> R2 (f x) (g y)" |
523 using q1 q2 r1 r2 a |
539 using q1 q2 r1 r2 a |
524 by (simp add: fun_rel_EQUALS) |
540 by (simp add: fun_rel_equals) |
525 |
541 |
526 lemma LAMBDA_REP_ABS_RSP: |
542 lemma lambda_rep_abs_rsp: |
527 assumes r1: "\<And>r r'. R1 r r' \<Longrightarrow>R1 r (Rep1 (Abs1 r'))" |
543 assumes r1: "\<And>r r'. R1 r r' \<Longrightarrow>R1 r (Rep1 (Abs1 r'))" |
528 and r2: "\<And>r r'. R2 r r' \<Longrightarrow>R2 r (Rep2 (Abs2 r'))" |
544 and r2: "\<And>r r'. R2 r r' \<Longrightarrow>R2 r (Rep2 (Abs2 r'))" |
529 shows "(R1 ===> R2) f1 f2 \<Longrightarrow> (R1 ===> R2) f1 ((Abs1 ---> Rep2) ((Rep1 ---> Abs2) f2))" |
545 shows "(R1 ===> R2) f1 f2 \<Longrightarrow> (R1 ===> R2) f1 ((Abs1 ---> Rep2) ((Rep1 ---> Abs2) f2))" |
530 using r1 r2 by auto |
546 using r1 r2 by auto |
531 |
|
532 (* Not used *) |
|
533 lemma rep_abs_rsp_left: |
|
534 assumes q: "Quotient R Abs Rep" |
|
535 and a: "R x1 x2" |
|
536 shows "R (Rep (Abs x1)) x2" |
|
537 using q a by (metis Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]) |
|
538 |
|
539 |
|
540 |
|
541 (* bool theory: COND, LET *) |
|
542 lemma IF_PRS: |
|
543 assumes q: "Quotient R Abs Rep" |
|
544 shows "If a b c = Abs (If a (Rep b) (Rep c))" |
|
545 using Quotient_abs_rep[OF q] by auto |
|
546 |
|
547 (* ask peter: no use of q *) |
|
548 lemma IF_RSP: |
|
549 assumes q: "Quotient R Abs Rep" |
|
550 and a: "a1 = a2" "R b1 b2" "R c1 c2" |
|
551 shows "R (If a1 b1 c1) (If a2 b2 c2)" |
|
552 using a by auto |
|
553 |
|
554 lemma LET_PRS: |
|
555 assumes q1: "Quotient R1 Abs1 Rep1" |
|
556 and q2: "Quotient R2 Abs2 Rep2" |
|
557 shows "Let x f = Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f))" |
|
558 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto |
|
559 |
|
560 lemma LET_RSP: |
|
561 assumes q1: "Quotient R1 Abs1 Rep1" |
|
562 and a1: "(R1 ===> R2) f g" |
|
563 and a2: "R1 x y" |
|
564 shows "R2 ((Let x f)::'c) ((Let y g)::'c)" |
|
565 using apply_rsp[OF q1 a1] a2 |
|
566 by auto |
|
567 |
|
568 |
|
569 |
547 |
570 (* ask peter what are literal_case *) |
548 (* ask peter what are literal_case *) |
571 (* literal_case_PRS *) |
549 (* literal_case_PRS *) |
572 (* literal_case_RSP *) |
550 (* literal_case_RSP *) |
573 |
551 (* Cez: !f x. literal_case f x = f x *) |
574 |
|
575 |
|
576 |
|
577 |
|
578 (* combinators: I, K, o, C, W *) |
|
579 |
552 |
580 (* We use id_simps which includes id_apply; so these 2 theorems can be removed *) |
553 (* We use id_simps which includes id_apply; so these 2 theorems can be removed *) |
581 |
554 lemma id_prs: |
582 lemma I_PRS: |
555 assumes q: "Quotient R Abs Rep" |
583 assumes q: "Quotient R Abs Rep" |
556 shows "Abs (id (Rep e)) = id e" |
584 shows "id e = Abs (id (Rep e))" |
557 using Quotient_abs_rep[OF q] by auto |
585 using Quotient_abs_rep[OF q] by auto |
558 |
586 |
559 lemma id_rsp: |
587 lemma I_RSP: |
|
588 assumes q: "Quotient R Abs Rep" |
560 assumes q: "Quotient R Abs Rep" |
589 and a: "R e1 e2" |
561 and a: "R e1 e2" |
590 shows "R (id e1) (id e2)" |
562 shows "R (id e1) (id e2)" |
591 using a by auto |
563 using a by auto |
592 |
|
593 lemma o_PRS: |
|
594 assumes q1: "Quotient R1 Abs1 Rep1" |
|
595 and q2: "Quotient R2 Abs2 Rep2" |
|
596 and q3: "Quotient R3 Abs3 Rep3" |
|
597 shows "f o g = (Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g))" |
|
598 using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3] |
|
599 unfolding o_def expand_fun_eq |
|
600 by simp |
|
601 |
|
602 lemma o_RSP: |
|
603 assumes q1: "Quotient R1 Abs1 Rep1" |
|
604 and q2: "Quotient R2 Abs2 Rep2" |
|
605 and q3: "Quotient R3 Abs3 Rep3" |
|
606 and a1: "(R2 ===> R3) f1 f2" |
|
607 and a2: "(R1 ===> R2) g1 g2" |
|
608 shows "(R1 ===> R3) (f1 o g1) (f2 o g2)" |
|
609 using a1 a2 unfolding o_def expand_fun_eq |
|
610 by (auto) |
|
611 |
|
612 lemma COND_PRS: |
|
613 assumes a: "Quotient R absf repf" |
|
614 shows "(if a then b else c) = absf (if a then repf b else repf c)" |
|
615 using a unfolding Quotient_def by auto |
|
616 |
|
617 |
564 |
618 end |
565 end |
619 |
566 |