src/HOL/Groups_List.thy
author haftmann
Sun, 30 Mar 2025 20:20:26 +0200
changeset 82388 f1ff9123c62a
parent 82097 25dd3726fd00
child 82597 328de89f20f9
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
     1
(* Author: Tobias Nipkow, TU Muenchen *)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
     2
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
     3
section \<open>Sum and product over lists\<close>
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
     4
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
     5
theory Groups_List
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
     6
imports List
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
     7
begin
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
     8
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
     9
locale monoid_list = monoid
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    10
begin
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    11
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    12
definition F :: "'a list \<Rightarrow> 'a"
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    13
where
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 63101
diff changeset
    14
  eq_foldr [code]: "F xs = foldr f xs \<^bold>1"
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    15
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    16
lemma Nil [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 63101
diff changeset
    17
  "F [] = \<^bold>1"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    18
  by (simp add: eq_foldr)
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    19
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    20
lemma Cons [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 63101
diff changeset
    21
  "F (x # xs) = x \<^bold>* F xs"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    22
  by (simp add: eq_foldr)
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    23
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    24
lemma append [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 63101
diff changeset
    25
  "F (xs @ ys) = F xs \<^bold>* F ys"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    26
  by (induct xs) (simp_all add: assoc)
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    27
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    28
end
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
    29
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    30
locale comm_monoid_list = comm_monoid + monoid_list
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    31
begin
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    32
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    33
lemma rev [simp]:
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    34
  "F (rev xs) = F xs"
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    35
  by (simp add: eq_foldr foldr_fold  fold_rev fun_eq_iff assoc left_commute)
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    36
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    37
end
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    38
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    39
locale comm_monoid_list_set = list: comm_monoid_list + set: comm_monoid_set
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    40
begin
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
    41
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    42
lemma distinct_set_conv_list:
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    43
  "distinct xs \<Longrightarrow> set.F g (set xs) = list.F (map g xs)"
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    44
  by (induct xs) simp_all
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
    45
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    46
lemma set_conv_list [code]:
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    47
  "set.F g (set xs) = list.F (map g (remdups xs))"
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    48
  by (simp add: distinct_set_conv_list [symmetric])
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    49
80061
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    50
lemma list_conv_set_nth:
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    51
  "list.F xs = set.F (\<lambda>i. xs ! i) {0..<length xs}"
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    52
proof -
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    53
  have "xs = map (\<lambda>i. xs ! i) [0..<length xs]"
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    54
    by (simp add: map_nth)
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    55
  also have "list.F \<dots> = set.F (\<lambda>i. xs ! i) {0..<length xs}"
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    56
    by (subst distinct_set_conv_list [symmetric]) auto
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    57
  finally show ?thesis .
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    58
qed
4c1347e172b1 moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
Manuel Eberl <eberlm@in.tum.de>
parents: 79017
diff changeset
    59
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    60
end
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    61
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    62
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
    63
subsection \<open>List summation\<close>
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    64
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    65
context monoid_add
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    66
begin
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    67
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    68
sublocale sum_list: monoid_list plus 0
61776
57bb7da5c867 modernized
haftmann
parents: 61605
diff changeset
    69
defines
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    70
  sum_list = sum_list.F ..
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
    71
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    72
end
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    73
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    74
context comm_monoid_add
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    75
begin
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    76
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    77
sublocale sum_list: comm_monoid_list plus 0
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61378
diff changeset
    78
rewrites
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    79
  "monoid_list.F plus 0 = sum_list"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    80
proof -
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    81
  show "comm_monoid_list plus 0" ..
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    82
  then interpret sum_list: comm_monoid_list plus 0 .
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    83
  from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
    84
qed
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
    85
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    86
sublocale sum: comm_monoid_list_set plus 0
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61378
diff changeset
    87
rewrites
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    88
  "monoid_list.F plus 0 = sum_list"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    89
  and "comm_monoid_set.F plus 0 = sum"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    90
proof -
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    91
  show "comm_monoid_list_set plus 0" ..
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    92
  then interpret sum: comm_monoid_list_set plus 0 .
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
    93
  from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
    94
  from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym)
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    95
qed
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    96
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    97
end
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
    98
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
    99
text \<open>Some syntactic sugar for summing a function over a list:\<close>
81595
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   100
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   101
open_bundle sum_list_syntax
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   102
begin
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   103
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61799
diff changeset
   104
syntax (ASCII)
80934
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   105
  "_sum_list" :: "pttrn => 'a list => 'b => 'b"    (\<open>(\<open>indent=3 notation=\<open>binder SUM\<close>\<close>SUM _<-_. _)\<close> [0, 51, 10] 10)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   106
syntax
80934
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   107
  "_sum_list" :: "pttrn => 'a list => 'b => 'b"    (\<open>(\<open>indent=3 notation=\<open>binder \<Sum>\<close>\<close>\<Sum>_\<leftarrow>_. _)\<close> [0, 51, 10] 10)
80760
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80061
diff changeset
   108
syntax_consts
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80061
diff changeset
   109
  "_sum_list" == sum_list
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61776
diff changeset
   110
translations \<comment> \<open>Beware of argument permutation!\<close>
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   111
  "\<Sum>x\<leftarrow>xs. b" == "CONST sum_list (CONST map (\<lambda>x. b) xs)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   112
81595
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   113
end
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   114
70928
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   115
context
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   116
  includes lifting_syntax
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   117
begin
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   118
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   119
lemma sum_list_transfer [transfer_rule]:
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   120
  "(list_all2 A ===> A) sum_list sum_list"
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   121
    if [transfer_rule]: "A 0 0" "(A ===> A ===> A) (+) (+)"
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   122
  unfolding sum_list.eq_foldr [abs_def]
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   123
  by transfer_prover
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   124
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   125
end
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   126
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
   127
text \<open>TODO duplicates\<close>
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   128
lemmas sum_list_simps = sum_list.Nil sum_list.Cons
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   129
lemmas sum_list_append = sum_list.append
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   130
lemmas sum_list_rev = sum_list.rev
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   131
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   132
lemma (in monoid_add) fold_plus_sum_list_rev:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   133
  "fold plus xs = plus (sum_list (rev xs))"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   134
proof
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   135
  fix x
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   136
  have "fold plus xs x = sum_list (rev xs @ [x])"
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   137
    by (simp add: foldr_conv_fold sum_list.eq_foldr)
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   138
  also have "\<dots> = sum_list (rev xs) + x"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   139
    by simp
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   140
  finally show "fold plus xs x = sum_list (rev xs) + x"
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   141
    .
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   142
qed
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   143
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   144
lemma sum_list_of_nat: "sum_list (map of_nat xs) = of_nat (sum_list xs)"
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   145
  by (induction xs) auto
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   146
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   147
lemma sum_list_of_int: "sum_list (map of_int xs) = of_int (sum_list xs)"
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   148
  by (induction xs) auto
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   149
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   150
lemma (in comm_monoid_add) sum_list_map_remove1:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   151
  "x \<in> set xs \<Longrightarrow> sum_list (map f xs) = f x + sum_list (map f (remove1 x xs))"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   152
  by (induct xs) (auto simp add: ac_simps)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   153
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   154
lemma (in monoid_add) size_list_conv_sum_list:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   155
  "size_list f xs = sum_list (map f xs) + size xs"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   156
  by (induct xs) auto
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   157
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   158
lemma (in monoid_add) length_concat:
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   159
  "length (concat xss) = sum_list (map length xss)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   160
  by (induct xss) simp_all
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   161
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   162
lemma (in monoid_add) length_product_lists:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67489
diff changeset
   163
  "length (product_lists xss) = foldr (*) (map length xss) 1"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   164
proof (induct xss)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   165
  case (Cons xs xss) then show ?case by (induct xs) (auto simp: length_concat o_def)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   166
qed simp
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   167
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   168
lemma (in monoid_add) sum_list_map_filter:
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   169
  assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> f x = 0"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   170
  shows "sum_list (map f (filter P xs)) = sum_list (map f xs)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   171
  using assms by (induct xs) auto
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   172
69231
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   173
lemma sum_list_filter_le_nat:
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   174
  fixes f :: "'a \<Rightarrow> nat"
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   175
  shows "sum_list (map f (filter P xs)) \<le> sum_list (map f xs)"
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   176
by(induction xs; simp)
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   177
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   178
lemma (in comm_monoid_add) distinct_sum_list_conv_Sum:
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   179
  "distinct xs \<Longrightarrow> sum_list xs = Sum (set xs)"
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   180
  by (metis local.sum.set_conv_list local.sum_list_def map_ident remdups_id_iff_distinct)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   181
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   182
lemma sum_list_upt[simp]:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   183
  "m \<le> n \<Longrightarrow> sum_list [m..<n] = \<Sum> {m..<n}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   184
by(simp add: distinct_sum_list_conv_Sum)
58995
42ba2b5cffac added lemma
nipkow
parents: 58889
diff changeset
   185
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   186
context ordered_comm_monoid_add
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   187
begin
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   188
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   189
lemma sum_list_nonneg: "(\<And>x. x \<in> set xs \<Longrightarrow> 0 \<le> x) \<Longrightarrow> 0 \<le> sum_list xs"
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   190
  by (induction xs) auto
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   191
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   192
lemma sum_list_nonpos: "(\<And>x. x \<in> set xs \<Longrightarrow> x \<le> 0) \<Longrightarrow> sum_list xs \<le> 0"
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   193
  by (induction xs) (auto simp: add_nonpos_nonpos)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   194
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   195
lemma sum_list_nonneg_eq_0_iff:
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   196
  "(\<And>x. x \<in> set xs \<Longrightarrow> 0 \<le> x) \<Longrightarrow> sum_list xs = 0 \<longleftrightarrow> (\<forall>x\<in> set xs. x = 0)"
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   197
  by (induction xs) (simp_all add: add_nonneg_eq_0_iff sum_list_nonneg)
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   198
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   199
end
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   200
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   201
context canonically_ordered_monoid_add
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   202
begin
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   203
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   204
lemma sum_list_eq_0_iff [simp]:
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   205
  "sum_list ns = 0 \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
82097
25dd3726fd00 Minor lemma tweaking
paulson <lp15@cam.ac.uk>
parents: 82080
diff changeset
   206
  by (simp add: sum_list_nonneg_eq_0_iff)
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   207
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   208
lemma member_le_sum_list:
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   209
  "x \<in> set xs \<Longrightarrow> x \<le> sum_list xs"
82097
25dd3726fd00 Minor lemma tweaking
paulson <lp15@cam.ac.uk>
parents: 82080
diff changeset
   210
  by (induction xs) (auto simp: add_increasing add_increasing2)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   211
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   212
lemma elem_le_sum_list:
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   213
  "k < size ns \<Longrightarrow> ns ! k \<le> sum_list (ns)"
82097
25dd3726fd00 Minor lemma tweaking
paulson <lp15@cam.ac.uk>
parents: 82080
diff changeset
   214
  by (simp add: member_le_sum_list)
66311
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   215
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   216
end
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   217
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   218
lemma (in ordered_cancel_comm_monoid_diff) sum_list_update:
037aaa0b6daf added lemmas
nipkow
parents: 66308
diff changeset
   219
  "k < size xs \<Longrightarrow> sum_list (xs[k := x]) = sum_list xs + x - xs ! k"
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   220
proof (induction xs arbitrary:k)
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   221
  case Nil
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   222
  then show ?case by auto
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   223
next
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   224
  case (Cons a xs)
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   225
  then show ?case
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   226
    apply (simp add: add_ac split: nat.split)
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   227
    using add_increasing diff_add_assoc elem_le_sum_list zero_le by force
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   228
qed
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   229
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   230
lemma (in monoid_add) sum_list_triv:
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   231
  "(\<Sum>x\<leftarrow>xs. r) = of_nat (length xs) * r"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   232
  by (induct xs) (simp_all add: distrib_right)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   233
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   234
lemma (in monoid_add) sum_list_0 [simp]:
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   235
  "(\<Sum>x\<leftarrow>xs. 0) = 0"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   236
  by (induct xs) (simp_all add: distrib_right)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   237
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61776
diff changeset
   238
text\<open>For non-Abelian groups \<open>xs\<close> needs to be reversed on one side:\<close>
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   239
lemma (in ab_group_add) uminus_sum_list_map:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   240
  "- sum_list (map f xs) = sum_list (map (uminus \<circ> f) xs)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   241
  by (induct xs) simp_all
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   242
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   243
lemma (in comm_monoid_add) sum_list_addf:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   244
  "(\<Sum>x\<leftarrow>xs. f x + g x) = sum_list (map f xs) + sum_list (map g xs)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   245
  by (induct xs) (simp_all add: algebra_simps)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   246
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   247
lemma (in ab_group_add) sum_list_subtractf:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   248
  "(\<Sum>x\<leftarrow>xs. f x - g x) = sum_list (map f xs) - sum_list (map g xs)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   249
  by (induct xs) (simp_all add: algebra_simps)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   250
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   251
lemma (in semiring_0) sum_list_const_mult:
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   252
  "(\<Sum>x\<leftarrow>xs. c * f x) = c * (\<Sum>x\<leftarrow>xs. f x)"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   253
  by (induct xs) (simp_all add: algebra_simps)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   254
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   255
lemma (in semiring_0) sum_list_mult_const:
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   256
  "(\<Sum>x\<leftarrow>xs. f x * c) = (\<Sum>x\<leftarrow>xs. f x) * c"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   257
  by (induct xs) (simp_all add: algebra_simps)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   258
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   259
lemma (in ordered_ab_group_add_abs) sum_list_abs:
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   260
  "\<bar>sum_list xs\<bar> \<le> sum_list (map abs xs)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   261
  by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq])
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   262
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   263
lemma sum_list_mono:
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   264
  fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, ordered_ab_semigroup_add}"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   265
  shows "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Sum>x\<leftarrow>xs. f x) \<le> (\<Sum>x\<leftarrow>xs. g x)"
69231
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   266
by (induct xs) (simp, simp add: add_mono)
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   267
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   268
lemma sum_list_strict_mono:
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   269
  fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, strict_ordered_ab_semigroup_add}"
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   270
  shows "\<lbrakk> xs \<noteq> [];  \<And>x. x \<in> set xs \<Longrightarrow> f x < g x \<rbrakk>
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   271
    \<Longrightarrow> sum_list (map f xs) < sum_list (map g xs)"
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   272
proof (induction xs)
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   273
  case Nil thus ?case by simp
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   274
next
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   275
  case C: (Cons _ xs)
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   276
  show ?case
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   277
  proof (cases xs)
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   278
    case Nil thus ?thesis using C.prems by simp
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   279
  next
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   280
    case Cons thus ?thesis using C by(simp add: add_strict_mono)
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   281
  qed
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   282
qed
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   283
75693
1d2222800ecd replaced complicated lemma by a simpler one
nipkow
parents: 75662
diff changeset
   284
text \<open>A much more general version of this monotonicity lemma
1d2222800ecd replaced complicated lemma by a simpler one
nipkow
parents: 75662
diff changeset
   285
can be formulated with multisets and the multiset order\<close>
1d2222800ecd replaced complicated lemma by a simpler one
nipkow
parents: 75662
diff changeset
   286
1d2222800ecd replaced complicated lemma by a simpler one
nipkow
parents: 75662
diff changeset
   287
lemma sum_list_mono2: fixes xs :: "'a ::ordered_comm_monoid_add list"
1d2222800ecd replaced complicated lemma by a simpler one
nipkow
parents: 75662
diff changeset
   288
shows "\<lbrakk> length xs = length ys; \<And>i. i < length xs \<longrightarrow> xs!i \<le> ys!i \<rbrakk>
1d2222800ecd replaced complicated lemma by a simpler one
nipkow
parents: 75662
diff changeset
   289
  \<Longrightarrow> sum_list xs \<le> sum_list ys"
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   290
  by (induction xs ys rule: list_induct2) (auto simp: nth_Cons' less_Suc_eq_0_disj imp_ex add_mono)
75693
1d2222800ecd replaced complicated lemma by a simpler one
nipkow
parents: 75662
diff changeset
   291
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   292
lemma (in monoid_add) sum_list_distinct_conv_sum_set:
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   293
  "distinct xs \<Longrightarrow> sum_list (map f xs) = sum f (set xs)"
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   294
  by (induct xs) simp_all
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   295
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   296
lemma (in monoid_add) interv_sum_list_conv_sum_set_nat:
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   297
  "sum_list (map f [m..<n]) = sum f (set [m..<n])"
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   298
  by (simp add: sum_list_distinct_conv_sum_set)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   299
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   300
lemma (in monoid_add) interv_sum_list_conv_sum_set_int:
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   301
  "sum_list (map f [k..l]) = sum f (set [k..l])"
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   302
  by (simp add: sum_list_distinct_conv_sum_set)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   303
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69231
diff changeset
   304
text \<open>General equivalence between \<^const>\<open>sum_list\<close> and \<^const>\<open>sum\<close>\<close>
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   305
lemma (in monoid_add) sum_list_sum_nth:
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   306
  "sum_list xs = (\<Sum> i = 0 ..< length xs. xs ! i)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66434
diff changeset
   307
  using interv_sum_list_conv_sum_set_nat [of "(!) xs" 0 "length xs"] by (simp add: map_nth)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   308
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   309
lemma sum_list_map_eq_sum_count:
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   310
  "sum_list (map f xs) = sum (\<lambda>x. count_list xs x * f x) (set xs)"
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   311
proof(induction xs)
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   312
  case (Cons x xs)
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   313
  show ?case (is "?l = ?r")
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   314
  proof cases
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   315
    assume "x \<in> set xs"
60541
4246da644cca modernized name
nipkow
parents: 59728
diff changeset
   316
    have "?l = f x + (\<Sum>x\<in>set xs. count_list xs x * f x)" by (simp add: Cons.IH)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
   317
    also have "set xs = insert x (set xs - {x})" using \<open>x \<in> set xs\<close>by blast
60541
4246da644cca modernized name
nipkow
parents: 59728
diff changeset
   318
    also have "f x + (\<Sum>x\<in>insert x (set xs - {x}). count_list xs x * f x) = ?r"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   319
      by (simp add: sum.insert_remove eq_commute)
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   320
    finally show ?thesis .
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   321
  next
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   322
    assume "x \<notin> set xs"
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   323
    hence "\<And>xa. xa \<in> set xs \<Longrightarrow> x \<noteq> xa" by blast
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
   324
    thus ?thesis by (simp add: Cons.IH \<open>x \<notin> set xs\<close>)
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   325
  qed
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   326
qed simp
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   327
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   328
lemma sum_list_map_eq_sum_count2:
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   329
assumes "set xs \<subseteq> X" "finite X"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   330
shows "sum_list (map f xs) = sum (\<lambda>x. count_list xs x * f x) X"
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   331
proof-
60541
4246da644cca modernized name
nipkow
parents: 59728
diff changeset
   332
  let ?F = "\<lambda>x. count_list xs x * f x"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   333
  have "sum ?F X = sum ?F (set xs \<union> (X - set xs))"
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   334
    using Un_absorb1[OF assms(1)] by(simp)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   335
  also have "\<dots> = sum ?F (set xs)"
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   336
    using assms(2)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   337
    by(simp add: sum.union_disjoint[OF _ _ Diff_disjoint] del: Un_Diff_cancel)
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   338
  finally show ?thesis by(simp add:sum_list_map_eq_sum_count)
59728
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   339
qed
0bb88aa34768 added lemmas
nipkow
parents: 58995
diff changeset
   340
72545
55a50f65c928 added lemma
nipkow
parents: 72187
diff changeset
   341
lemma sum_list_replicate: "sum_list (replicate n c) = of_nat n * c"
55a50f65c928 added lemma
nipkow
parents: 72187
diff changeset
   342
by(induction n)(auto simp add: distrib_right)
55a50f65c928 added lemma
nipkow
parents: 72187
diff changeset
   343
55a50f65c928 added lemma
nipkow
parents: 72187
diff changeset
   344
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   345
lemma sum_list_nonneg:
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   346
    "(\<And>x. x \<in> set xs \<Longrightarrow> (x :: 'a :: ordered_comm_monoid_add) \<ge> 0) \<Longrightarrow> sum_list xs \<ge> 0"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 61955
diff changeset
   347
  by (induction xs) simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 61955
diff changeset
   348
69231
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   349
lemma sum_list_Suc:
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   350
  "sum_list (map (\<lambda>x. Suc(f x)) xs) = sum_list (map f xs) + length xs"
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   351
by(induction xs; simp)
6b90ace5e5eb more lemmas
nipkow
parents: 69064
diff changeset
   352
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   353
lemma (in monoid_add) sum_list_map_filter':
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   354
  "sum_list (map f (filter P xs)) = sum_list (map (\<lambda>x. if P x then f x else 0) xs)"
63099
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 61955
diff changeset
   355
  by (induction xs) simp_all
af0e964aad7b Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents: 61955
diff changeset
   356
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   357
text \<open>Summation of a strictly ascending sequence with length \<open>n\<close>
66434
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   358
  can be upper-bounded by summation over \<open>{0..<n}\<close>.\<close>
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   359
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   360
lemma sorted_wrt_less_sum_mono_lowerbound:
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   361
  fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   362
  assumes mono: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66434
diff changeset
   363
  shows "sorted_wrt (<) ns \<Longrightarrow>
66434
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   364
    (\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   365
proof (induction ns rule: rev_induct)
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   366
  case Nil
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   367
  then show ?case by simp
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   368
next
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   369
  case (snoc n ns)
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   370
  have "sum f {0..<length (ns @ [n])}
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   371
      = sum f {0..<length ns} + f (length ns)"
66434
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   372
    by simp
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   373
  also have "sum f {0..<length ns} \<le> sum_list (map f ns)"
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   374
    using snoc by (auto simp: sorted_wrt_append)
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   375
  also have "length ns \<le> n"
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   376
    using sorted_wrt_less_idx[OF snoc.prems(1), of "length ns"] by auto
66434
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   377
  finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   378
    using mono add_mono by blast
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   379
  thus ?case by simp
67489
f1ba59ddd9a6 drop redundant cong rules
Lars Hupel <lars.hupel@mytum.de>
parents: 67399
diff changeset
   380
qed
66434
5d7e770c7d5d added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
nipkow
parents: 66311
diff changeset
   381
82097
25dd3726fd00 Minor lemma tweaking
paulson <lp15@cam.ac.uk>
parents: 82080
diff changeset
   382
(*Note that we also have this for class canonically_ordered_monoid_add*)
82080
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   383
lemma member_le_sum_list:
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   384
  fixes x :: "'a :: ordered_comm_monoid_add"
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   385
  assumes "x \<in> set xs" "\<And>x. x \<in> set xs \<Longrightarrow> x \<ge> 0"
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   386
  shows   "x \<le> sum_list xs"
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   387
  using assms
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   388
proof (induction xs)
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   389
  case (Cons y xs)
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   390
  show ?case
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   391
  proof (cases "y = x")
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   392
    case True
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   393
    have "x + 0 \<le> x + sum_list xs"
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   394
      by (intro add_mono order.refl sum_list_nonneg) (use Cons in auto)
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   395
    thus ?thesis
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   396
      using True by auto
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   397
  next
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   398
    case False
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   399
    have "0 + x \<le> y + sum_list xs"
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   400
      by (intro add_mono Cons.IH Cons.prems) (use Cons.prems False in auto)
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   401
    thus ?thesis
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   402
      by auto
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   403
  qed
0aa2d1c132b2 A couple of theorems proved by Manuel Eberl from his AFP entry Sum_Of_Squares_Count
paulson <lp15@cam.ac.uk>
parents: 81595
diff changeset
   404
qed auto
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   405
72024
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   406
subsection \<open>Horner sums\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   407
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   408
context comm_semiring_0
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   409
begin
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   410
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   411
definition horner_sum :: \<open>('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'a\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   412
  where horner_sum_foldr: \<open>horner_sum f a xs = foldr (\<lambda>x b. f x + a * b) xs 0\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   413
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   414
lemma horner_sum_simps [simp]:
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   415
  \<open>horner_sum f a [] = 0\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   416
  \<open>horner_sum f a (x # xs) = f x + a * horner_sum f a xs\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   417
  by (simp_all add: horner_sum_foldr)
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   418
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   419
lemma horner_sum_eq_sum_funpow:
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   420
  \<open>horner_sum f a xs = (\<Sum>n = 0..<length xs. ((*) a ^^ n) (f (xs ! n)))\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   421
proof (induction xs)
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   422
  case Nil
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   423
  then show ?case
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   424
    by simp
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   425
next
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   426
  case (Cons x xs)
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   427
  then show ?case
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   428
    by (simp add: sum.atLeast0_lessThan_Suc_shift sum_distrib_left del: sum.op_ivl_Suc)
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   429
qed
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   430
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   431
end
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   432
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   433
context
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   434
  includes lifting_syntax
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   435
begin
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   436
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   437
lemma horner_sum_transfer [transfer_rule]:
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   438
  \<open>((B ===> A) ===> A ===> list_all2 B ===> A) horner_sum horner_sum\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   439
  if [transfer_rule]: \<open>A 0 0\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   440
    and [transfer_rule]: \<open>(A ===> A ===> A) (+) (+)\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   441
    and [transfer_rule]: \<open>(A ===> A ===> A) (*) (*)\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   442
  by (unfold horner_sum_foldr) transfer_prover
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   443
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   444
end
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   445
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   446
context comm_semiring_1
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   447
begin
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   448
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   449
lemma horner_sum_eq_sum:
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   450
  \<open>horner_sum f a xs = (\<Sum>n = 0..<length xs. f (xs ! n) * a ^ n)\<close>
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   451
proof -
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   452
  have \<open>(*) a ^^ n = (*) (a ^ n)\<close> for n
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   453
    by (induction n) (simp_all add: ac_simps)
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   454
  then show ?thesis
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   455
    by (simp add: horner_sum_eq_sum_funpow ac_simps)
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   456
qed
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   457
72619
4b2691211719 moved lemmas from AFP to distribution
haftmann
parents: 72545
diff changeset
   458
lemma horner_sum_append:
4b2691211719 moved lemmas from AFP to distribution
haftmann
parents: 72545
diff changeset
   459
  \<open>horner_sum f a (xs @ ys) = horner_sum f a xs + a ^ length xs * horner_sum f a ys\<close>
4b2691211719 moved lemmas from AFP to distribution
haftmann
parents: 72545
diff changeset
   460
  using sum.atLeastLessThan_shift_bounds [of _ 0 \<open>length xs\<close> \<open>length ys\<close>]
4b2691211719 moved lemmas from AFP to distribution
haftmann
parents: 72545
diff changeset
   461
    atLeastLessThan_add_Un [of 0 \<open>length xs\<close> \<open>length ys\<close>]
4b2691211719 moved lemmas from AFP to distribution
haftmann
parents: 72545
diff changeset
   462
  by (simp add: horner_sum_eq_sum sum_distrib_left sum.union_disjoint ac_simps nth_append power_add)
4b2691211719 moved lemmas from AFP to distribution
haftmann
parents: 72545
diff changeset
   463
72024
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   464
end
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   465
75662
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   466
context linordered_semidom
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   467
begin
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   468
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   469
lemma horner_sum_nonnegative:
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   470
  \<open>0 \<le> horner_sum of_bool 2 bs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   471
  by (induction bs) simp_all
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   472
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   473
end
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   474
78935
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   475
context discrete_linordered_semidom
75662
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   476
begin
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   477
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   478
lemma horner_sum_bound:
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   479
  \<open>horner_sum of_bool 2 bs < 2 ^ length bs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   480
proof (induction bs)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   481
  case Nil
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   482
  then show ?case
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   483
    by simp
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   484
next
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   485
  case (Cons b bs)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   486
  moreover define a where \<open>a = 2 ^ length bs - horner_sum of_bool 2 bs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   487
  ultimately have *: \<open>2 ^ length bs = horner_sum of_bool 2 bs + a\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   488
    by simp
78935
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   489
  have \<open>0 < a\<close>
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   490
    using Cons * by simp
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   491
  moreover have \<open>1 \<le> a\<close>
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   492
    using \<open>0 < a\<close> by (simp add: less_eq_iff_succ_less)
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   493
  ultimately have \<open>0 + 1 < a + a\<close>
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   494
    by (rule add_less_le_mono)
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   495
  then have \<open>1 < a * 2\<close>
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   496
    by (simp add: mult_2_right)
75662
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   497
  with Cons show ?case
78935
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   498
    by (simp add: * algebra_simps)
75662
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   499
qed
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   500
79017
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 78935
diff changeset
   501
lemma horner_sum_of_bool_2_less:
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 78935
diff changeset
   502
  \<open>(horner_sum of_bool 2 bs) < 2 ^ length bs\<close>
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 78935
diff changeset
   503
  by (fact horner_sum_bound)
127ba61b2630 modernized, reordered, generalized
haftmann
parents: 78935
diff changeset
   504
75662
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   505
end
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   506
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   507
lemma nat_horner_sum [simp]:
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   508
  \<open>nat (horner_sum of_bool 2 bs) = horner_sum of_bool 2 bs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   509
  by (induction bs) (auto simp add: nat_add_distrib horner_sum_nonnegative)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   510
78935
5e788ff7a489 explicit type class for discrete linordered semidoms
haftmann
parents: 75693
diff changeset
   511
context discrete_linordered_semidom
75662
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   512
begin
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   513
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   514
lemma horner_sum_less_eq_iff_lexordp_eq:
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   515
  \<open>horner_sum of_bool 2 bs \<le> horner_sum of_bool 2 cs \<longleftrightarrow> lexordp_eq (rev bs) (rev cs)\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   516
  if \<open>length bs = length cs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   517
proof -
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   518
  have \<open>horner_sum of_bool 2 (rev bs) \<le> horner_sum of_bool 2 (rev cs) \<longleftrightarrow> lexordp_eq bs cs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   519
    if \<open>length bs = length cs\<close> for bs cs
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   520
  using that proof (induction bs cs rule: list_induct2)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   521
    case Nil
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   522
    then show ?case
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   523
      by simp
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   524
  next
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   525
    case (Cons b bs c cs)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   526
    with horner_sum_nonnegative [of \<open>rev bs\<close>] horner_sum_nonnegative [of \<open>rev cs\<close>]
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   527
      horner_sum_bound [of \<open>rev bs\<close>] horner_sum_bound [of \<open>rev cs\<close>]
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   528
    show ?case
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   529
      by (auto simp add: horner_sum_append not_le Cons intro: add_strict_increasing2 add_increasing)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   530
  qed
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   531
  from that this [of \<open>rev bs\<close> \<open>rev cs\<close>] show ?thesis
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   532
    by simp
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   533
qed
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   534
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   535
lemma horner_sum_less_iff_lexordp:
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   536
  \<open>horner_sum of_bool 2 bs < horner_sum of_bool 2 cs \<longleftrightarrow> ord_class.lexordp (rev bs) (rev cs)\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   537
  if \<open>length bs = length cs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   538
proof -
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   539
  have \<open>horner_sum of_bool 2 (rev bs) < horner_sum of_bool 2 (rev cs) \<longleftrightarrow> ord_class.lexordp bs cs\<close>
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   540
    if \<open>length bs = length cs\<close> for bs cs
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   541
  using that proof (induction bs cs rule: list_induct2)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   542
    case Nil
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   543
    then show ?case
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   544
      by simp
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   545
  next
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   546
    case (Cons b bs c cs)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   547
    with horner_sum_nonnegative [of \<open>rev bs\<close>] horner_sum_nonnegative [of \<open>rev cs\<close>]
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   548
      horner_sum_bound [of \<open>rev bs\<close>] horner_sum_bound [of \<open>rev cs\<close>]
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   549
    show ?case
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   550
      by (auto simp add: horner_sum_append not_less Cons intro: add_strict_increasing2 add_increasing)
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   551
  qed
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   552
  from that this [of \<open>rev bs\<close> \<open>rev cs\<close>] show ?thesis
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   553
    by simp
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   554
qed
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   555
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   556
end
ed15f2cd4f7d refined code equations for characters
haftmann
parents: 74101
diff changeset
   557
72024
9b4135e8bade a generic horner sum operation
haftmann
parents: 70928
diff changeset
   558
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69231
diff changeset
   559
subsection \<open>Further facts about \<^const>\<open>List.n_lists\<close>\<close>
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   560
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   561
lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   562
  by (induct n) (auto simp add: comp_def length_concat sum_list_triv)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   563
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   564
lemma distinct_n_lists:
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   565
  assumes "distinct xs"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   566
  shows "distinct (List.n_lists n xs)"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   567
proof (rule card_distinct)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   568
  from assms have card_length: "card (set xs) = length xs" by (rule distinct_card)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   569
  have "card (set (List.n_lists n xs)) = card (set xs) ^ n"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   570
  proof (induct n)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   571
    case 0 then show ?case by simp
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   572
  next
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   573
    case (Suc n)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   574
    moreover have "card (\<Union>ys\<in>set (List.n_lists n xs). (\<lambda>y. y # ys) ` set xs)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   575
      = (\<Sum>ys\<in>set (List.n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   576
      by (rule card_UN_disjoint) auto
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   577
    moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   578
      by (rule card_image) (simp add: inj_on_def)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   579
    ultimately show ?case by auto
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   580
  qed
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   581
  also have "\<dots> = length xs ^ n" by (simp add: card_length)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   582
  finally show "card (set (List.n_lists n xs)) = length (List.n_lists n xs)"
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   583
    by (simp add: length_n_lists)
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   584
qed
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   585
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   586
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
   587
subsection \<open>Tools setup\<close>
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   588
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   589
lemmas sum_code = sum.set_conv_list
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   590
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   591
lemma sum_set_upto_conv_sum_list_int [code_unfold]:
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   592
  "sum f (set [i..j::int]) = sum_list (map f [i..j])"
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   593
  by (simp add: interv_sum_list_conv_sum_set_int)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   594
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   595
lemma sum_set_upt_conv_sum_list_nat [code_unfold]:
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   596
  "sum f (set [m..<n]) = sum_list (map f [m..<n])"
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   597
  by (simp add: interv_sum_list_conv_sum_set_nat)
58101
e7ebe5554281 separated listsum material
haftmann
parents:
diff changeset
   598
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   599
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
   600
subsection \<open>List product\<close>
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   601
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   602
context monoid_mult
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   603
begin
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   604
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   605
sublocale prod_list: monoid_list times 1
61776
57bb7da5c867 modernized
haftmann
parents: 61605
diff changeset
   606
defines
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   607
  prod_list = prod_list.F ..
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   608
58320
351810c45a48 abstract product over monoid for lists
haftmann
parents: 58152
diff changeset
   609
end
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   610
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   611
context comm_monoid_mult
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   612
begin
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   613
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   614
sublocale prod_list: comm_monoid_list times 1
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61378
diff changeset
   615
rewrites
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   616
  "monoid_list.F times 1 = prod_list"
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   617
proof -
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   618
  show "comm_monoid_list times 1" ..
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   619
  then interpret prod_list: comm_monoid_list times 1 .
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   620
  from prod_list_def show "monoid_list.F times 1 = prod_list" by simp
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   621
qed
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   622
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   623
sublocale prod: comm_monoid_list_set times 1
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61378
diff changeset
   624
rewrites
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   625
  "monoid_list.F times 1 = prod_list"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   626
  and "comm_monoid_set.F times 1 = prod"
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   627
proof -
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   628
  show "comm_monoid_list_set times 1" ..
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   629
  then interpret prod: comm_monoid_list_set times 1 .
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   630
  from prod_list_def show "monoid_list.F times 1 = prod_list" by simp
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   631
  from prod_def show "comm_monoid_set.F times 1 = prod" by (auto intro: sym)
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   632
qed
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   633
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   634
end
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   635
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60541
diff changeset
   636
text \<open>Some syntactic sugar:\<close>
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   637
81595
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   638
open_bundle prod_list_syntax
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   639
begin
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   640
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61799
diff changeset
   641
syntax (ASCII)
80934
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   642
  "_prod_list" :: "pttrn => 'a list => 'b => 'b"    (\<open>(\<open>indent=3 notation=\<open>binder PROD\<close>\<close>PROD _<-_. _)\<close> [0, 51, 10] 10)
61955
e96292f32c3c former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents: 61799
diff changeset
   643
syntax
80934
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   644
  "_prod_list" :: "pttrn => 'a list => 'b => 'b"    (\<open>(\<open>indent=3 notation=\<open>binder \<Prod>\<close>\<close>\<Prod>_\<leftarrow>_. _)\<close> [0, 51, 10] 10)
80760
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80061
diff changeset
   645
syntax_consts
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80061
diff changeset
   646
  "_prod_list" \<rightleftharpoons> prod_list
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61776
diff changeset
   647
translations \<comment> \<open>Beware of argument permutation!\<close>
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63343
diff changeset
   648
  "\<Prod>x\<leftarrow>xs. b" \<rightleftharpoons> "CONST prod_list (CONST map (\<lambda>x. b) xs)"
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   649
81595
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   650
end
ed264056f5dc more syntax bundles, e.g. to explore terms without notation;
wenzelm
parents: 80934
diff changeset
   651
70928
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   652
context
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   653
  includes lifting_syntax
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   654
begin
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   655
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   656
lemma prod_list_transfer [transfer_rule]:
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   657
  "(list_all2 A ===> A) prod_list prod_list"
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   658
    if [transfer_rule]: "A 1 1" "(A ===> A ===> A) (*) (*)"
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   659
  unfolding prod_list.eq_foldr [abs_def]
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   660
  by transfer_prover
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   661
58368
fe083c681ed8 product over monoids for lists
haftmann
parents: 58320
diff changeset
   662
end
70928
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   663
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   664
lemma prod_list_zero_iff:
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   665
  "prod_list xs = 0 \<longleftrightarrow> (0 :: 'a :: {semiring_no_zero_divisors, semiring_1}) \<in> set xs"
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   666
  by (induction xs) simp_all
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   667
273fc913427b more transfer rules
haftmann
parents: 70927
diff changeset
   668
end