Quot/QuotProd.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 09 Dec 2009 17:05:33 +0100
changeset 668 ef5b941f00e2
parent 648 830b58c2fa94
child 699 aa157e957655
permissions -rw-r--r--
Code cleaning.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory QuotProd
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
     2
imports QuotMain
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
fun
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
  prod_rel
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
where
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
  "prod_rel r1 r2 = (\<lambda>(a,b) (c,d). r1 a c \<and> r2 b d)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
(* prod_fun is a good mapping function *)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
lemma prod_equivp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
  assumes a: "equivp R1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
  assumes b: "equivp R2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
  shows "equivp (prod_rel R1 R2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
apply(auto simp add: equivp_reflp[OF a] equivp_reflp[OF b])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
apply(simp only: equivp_symp[OF a])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
apply(simp only: equivp_symp[OF b])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
using equivp_transp[OF a] apply blast
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
using equivp_transp[OF b] apply blast
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
done
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
lemma prod_quotient:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
  shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
unfolding Quotient_def
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
apply (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q1] Quotient_rel_rep[OF q2])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
using Quotient_rel[OF q1] Quotient_rel[OF q2] by blast
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
lemma pair_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
  shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
by auto
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    38
lemma pair_prs_aux:
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
  shows "(prod_fun Abs1 Abs2) (Rep1 l, Rep2 r) \<equiv> (l, r)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
  by (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    44
term Pair
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    45
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    46
lemma pair_prs[quot_preserve]:
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    47
  assumes q1: "Quotient R1 Abs1 Rep1"
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    48
  assumes q2: "Quotient R2 Abs2 Rep2"
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    49
  shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair"
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    50
apply(simp only: expand_fun_eq fun_map.simps pair_prs_aux[OF q1 q2])
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    51
apply(simp)
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    52
done
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    53
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    54
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    55
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    56
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
(* TODO: Is the quotient assumption q1 necessary? *)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
(* TODO: Aren't there hard to use later? *)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
lemma fst_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
  assumes a: "(prod_rel R1 R2) p1 p2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
  shows "R1 (fst p1) (fst p2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
  using a
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
  apply(case_tac p1)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
  apply(case_tac p2)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
  apply(auto)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
  done
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
lemma snd_rsp:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
  assumes a: "(prod_rel R1 R2) p1 p2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
  shows "R2 (snd p1) (snd p2)"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
  using a
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
  apply(case_tac p1)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
  apply(case_tac p2)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
  apply(auto)
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
  done
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
lemma fst_prs:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
  shows "Abs1 (fst ((prod_fun Rep1 Rep2) p)) = fst p"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
by (case_tac p) (auto simp add: Quotient_abs_rep[OF q1])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
lemma snd_prs:
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
  assumes q1: "Quotient R1 Abs1 Rep1"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
  assumes q2: "Quotient R2 Abs2 Rep2"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
  shows "Abs2 (snd ((prod_fun Rep1 Rep2) p)) = snd p"
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
by (case_tac p) (auto simp add: Quotient_abs_rep[OF q2])
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
668
ef5b941f00e2 Code cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 648
diff changeset
    93
lemma prod_fun_id[id_simps]: "prod_fun id id \<equiv> id"
ef5b941f00e2 Code cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 648
diff changeset
    94
  by (rule eq_reflection) (simp add: prod_fun_def)
648
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    95
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    96
section {* general setup for products *}
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    97
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    98
declare [[map * = (prod_fun, prod_rel)]]
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
    99
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   100
lemmas [quot_thm] = prod_quotient
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   101
lemmas [quot_respect] = pair_rsp
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   102
lemmas [quot_equiv] = prod_equivp
830b58c2fa94 decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
parents: 597
diff changeset
   103
593
18eac4596ef1 QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
end