src/HOL/Library/Product_Plus.thy
author wenzelm
Wed, 08 Mar 2017 10:50:59 +0100
changeset 65151 a7394aa4d21c
parent 64267 b9a1486e79be
permissions -rw-r--r--
tuned proofs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63972
c98d1dd7eba1 Library: fix name Product_plus to Product_Plus
hoelzl
parents: 60679
diff changeset
     1
(*  Title:      HOL/Library/Product_Plus.thy
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
     3
*)
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
     4
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
     5
section \<open>Additive group operations on product types\<close>
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
     6
63972
c98d1dd7eba1 Library: fix name Product_plus to Product_Plus
hoelzl
parents: 60679
diff changeset
     7
theory Product_Plus
51301
6822aa82aafa simplified imports;
wenzelm
parents: 51002
diff changeset
     8
imports Main
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
     9
begin
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    10
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
    11
subsection \<open>Operations\<close>
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    12
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    13
instantiation prod :: (zero, zero) zero
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    14
begin
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    15
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    16
definition zero_prod_def: "0 = (0, 0)"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    17
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    18
instance ..
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    19
end
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    20
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    21
instantiation prod :: (plus, plus) plus
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    22
begin
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    23
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    24
definition plus_prod_def:
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    25
  "x + y = (fst x + fst y, snd x + snd y)"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    26
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    27
instance ..
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    28
end
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    29
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    30
instantiation prod :: (minus, minus) minus
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    31
begin
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    32
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    33
definition minus_prod_def:
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    34
  "x - y = (fst x - fst y, snd x - snd y)"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    35
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    36
instance ..
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    37
end
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    38
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    39
instantiation prod :: (uminus, uminus) uminus
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    40
begin
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    41
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    42
definition uminus_prod_def:
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    43
  "- x = (- fst x, - snd x)"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    44
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    45
instance ..
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    46
end
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    47
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    48
lemma fst_zero [simp]: "fst 0 = 0"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    49
  unfolding zero_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    50
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    51
lemma snd_zero [simp]: "snd 0 = 0"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    52
  unfolding zero_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    53
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    54
lemma fst_add [simp]: "fst (x + y) = fst x + fst y"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    55
  unfolding plus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    56
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    57
lemma snd_add [simp]: "snd (x + y) = snd x + snd y"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    58
  unfolding plus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    59
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    60
lemma fst_diff [simp]: "fst (x - y) = fst x - fst y"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    61
  unfolding minus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    62
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    63
lemma snd_diff [simp]: "snd (x - y) = snd x - snd y"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    64
  unfolding minus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    65
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    66
lemma fst_uminus [simp]: "fst (- x) = - fst x"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    67
  unfolding uminus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    68
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    69
lemma snd_uminus [simp]: "snd (- x) = - snd x"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    70
  unfolding uminus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    71
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    72
lemma add_Pair [simp]: "(a, b) + (c, d) = (a + c, b + d)"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    73
  unfolding plus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    74
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    75
lemma diff_Pair [simp]: "(a, b) - (c, d) = (a - c, b - d)"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    76
  unfolding minus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    77
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    78
lemma uminus_Pair [simp, code]: "- (a, b) = (- a, - b)"
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    79
  unfolding uminus_prod_def by simp
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    80
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59815
diff changeset
    81
subsection \<open>Class instances\<close>
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    82
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    83
instance prod :: (semigroup_add, semigroup_add) semigroup_add
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    84
  by standard (simp add: prod_eq_iff add.assoc)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    85
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    86
instance prod :: (ab_semigroup_add, ab_semigroup_add) ab_semigroup_add
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    87
  by standard (simp add: prod_eq_iff add.commute)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    88
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    89
instance prod :: (monoid_add, monoid_add) monoid_add
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    90
  by standard (simp_all add: prod_eq_iff)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    91
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
    92
instance prod :: (comm_monoid_add, comm_monoid_add) comm_monoid_add
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    93
  by standard (simp add: prod_eq_iff)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    94
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    95
instance prod :: (cancel_semigroup_add, cancel_semigroup_add) cancel_semigroup_add
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    96
    by standard (simp_all add: prod_eq_iff)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
    97
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    98
instance prod :: (cancel_ab_semigroup_add, cancel_ab_semigroup_add) cancel_ab_semigroup_add
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
    99
    by standard (simp_all add: prod_eq_iff diff_diff_eq)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
   100
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   101
instance prod :: (cancel_comm_monoid_add, cancel_comm_monoid_add) cancel_comm_monoid_add ..
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
   102
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
   103
instance prod :: (group_add, group_add) group_add
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   104
  by standard (simp_all add: prod_eq_iff)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
   105
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37664
diff changeset
   106
instance prod :: (ab_group_add, ab_group_add) ab_group_add
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   107
  by standard (simp_all add: prod_eq_iff)
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
   108
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63972
diff changeset
   109
lemma fst_sum: "fst (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. fst (f x))"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   110
proof (cases "finite A")
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   111
  case True
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   112
  then show ?thesis by induct simp_all
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   113
next
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   114
  case False
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   115
  then show ?thesis by simp
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   116
qed
36627
39b2516a1970 move setsum lemmas to Product_plus.thy
huffman
parents: 30018
diff changeset
   117
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63972
diff changeset
   118
lemma snd_sum: "snd (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. snd (f x))"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   119
proof (cases "finite A")
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   120
  case True
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   121
  then show ?thesis by induct simp_all
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   122
next
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   123
  case False
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   124
  then show ?thesis by simp
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   125
qed
36627
39b2516a1970 move setsum lemmas to Product_plus.thy
huffman
parents: 30018
diff changeset
   126
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63972
diff changeset
   127
lemma sum_prod: "(\<Sum>x\<in>A. (f x, g x)) = (\<Sum>x\<in>A. f x, \<Sum>x\<in>A. g x)"
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   128
proof (cases "finite A")
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   129
  case True
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   130
  then show ?thesis by induct (simp_all add: zero_prod_def)
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   131
next
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   132
  case False
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   133
  then show ?thesis by (simp add: zero_prod_def)
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   134
qed
37664
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 36627
diff changeset
   135
30018
690c65b8ad1a add new theory Product_plus.thy to Library
huffman
parents:
diff changeset
   136
end