src/HOL/Algebra/FiniteProduct.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 15328 35951e6a7855
child 16638 3dc904d93767
permissions -rw-r--r--
migrated theory headers to new format
wenzelm@14706
     1
(*  Title:      HOL/Algebra/FiniteProduct.thy
ballarin@13936
     2
    ID:         $Id$
ballarin@13936
     3
    Author:     Clemens Ballarin, started 19 November 2002
ballarin@13936
     4
ballarin@13936
     5
This file is largely based on HOL/Finite_Set.thy.
ballarin@13936
     6
*)
ballarin@13936
     7
wenzelm@14706
     8
header {* Product Operator for Commutative Monoids *}
ballarin@13936
     9
haftmann@16417
    10
theory FiniteProduct imports Group begin
ballarin@13936
    11
paulson@14750
    12
text {* Instantiation of locale @{text LC} of theory @{text Finite_Set} is not
paulson@14750
    13
  possible, because here we have explicit typing rules like 
paulson@14750
    14
  @{text "x \<in> carrier G"}.  We introduce an explicit argument for the domain
wenzelm@14651
    15
  @{text D}. *}
ballarin@13936
    16
ballarin@13936
    17
consts
ballarin@13936
    18
  foldSetD :: "['a set, 'b => 'a => 'a, 'a] => ('b set * 'a) set"
ballarin@13936
    19
ballarin@13936
    20
inductive "foldSetD D f e"
ballarin@13936
    21
  intros
paulson@14750
    22
    emptyI [intro]: "e \<in> D ==> ({}, e) \<in> foldSetD D f e"
paulson@14750
    23
    insertI [intro]: "[| x ~: A; f x y \<in> D; (A, y) \<in> foldSetD D f e |] ==>
paulson@14750
    24
                      (insert x A, f x y) \<in> foldSetD D f e"
ballarin@13936
    25
paulson@14750
    26
inductive_cases empty_foldSetDE [elim!]: "({}, x) \<in> foldSetD D f e"
ballarin@13936
    27
ballarin@13936
    28
constdefs
ballarin@13936
    29
  foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a"
paulson@14750
    30
  "foldD D f e A == THE x. (A, x) \<in> foldSetD D f e"
ballarin@13936
    31
ballarin@13936
    32
lemma foldSetD_closed:
paulson@14750
    33
  "[| (A, z) \<in> foldSetD D f e ; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D 
paulson@14750
    34
      |] ==> z \<in> D";
ballarin@13936
    35
  by (erule foldSetD.elims) auto
ballarin@13936
    36
ballarin@13936
    37
lemma Diff1_foldSetD:
paulson@14750
    38
  "[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; f x y \<in> D |] ==>
paulson@14750
    39
   (A, f x y) \<in> foldSetD D f e"
ballarin@13936
    40
  apply (erule insert_Diff [THEN subst], rule foldSetD.intros)
ballarin@13936
    41
    apply auto
ballarin@13936
    42
  done
ballarin@13936
    43
paulson@14750
    44
lemma foldSetD_imp_finite [simp]: "(A, x) \<in> foldSetD D f e ==> finite A"
ballarin@13936
    45
  by (induct set: foldSetD) auto
ballarin@13936
    46
ballarin@13936
    47
lemma finite_imp_foldSetD:
paulson@14750
    48
  "[| finite A; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D |] ==>
paulson@14750
    49
   EX x. (A, x) \<in> foldSetD D f e"
ballarin@13936
    50
proof (induct set: Finites)
ballarin@13936
    51
  case empty then show ?case by auto
ballarin@13936
    52
next
nipkow@15328
    53
  case (insert x F)
paulson@14750
    54
  then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto
paulson@14750
    55
  with insert have "y \<in> D" by (auto dest: foldSetD_closed)
paulson@14750
    56
  with y and insert have "(insert x F, f x y) \<in> foldSetD D f e"
ballarin@13936
    57
    by (intro foldSetD.intros) auto
ballarin@13936
    58
  then show ?case ..
ballarin@13936
    59
qed
ballarin@13936
    60
ballarin@13936
    61
subsection {* Left-commutative operations *}
ballarin@13936
    62
ballarin@13936
    63
locale LCD =
ballarin@13936
    64
  fixes B :: "'b set"
ballarin@13936
    65
  and D :: "'a set"
ballarin@13936
    66
  and f :: "'b => 'a => 'a"    (infixl "\<cdot>" 70)
ballarin@13936
    67
  assumes left_commute:
paulson@14750
    68
    "[| x \<in> B; y \<in> B; z \<in> D |] ==> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
paulson@14750
    69
  and f_closed [simp, intro!]: "!!x y. [| x \<in> B; y \<in> D |] ==> f x y \<in> D"
ballarin@13936
    70
ballarin@13936
    71
lemma (in LCD) foldSetD_closed [dest]:
paulson@14750
    72
  "(A, z) \<in> foldSetD D f e ==> z \<in> D";
ballarin@13936
    73
  by (erule foldSetD.elims) auto
ballarin@13936
    74
ballarin@13936
    75
lemma (in LCD) Diff1_foldSetD:
paulson@14750
    76
  "[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; A \<subseteq> B |] ==>
paulson@14750
    77
  (A, f x y) \<in> foldSetD D f e"
paulson@14750
    78
  apply (subgoal_tac "x \<in> B")
ballarin@13936
    79
   prefer 2 apply fast
ballarin@13936
    80
  apply (erule insert_Diff [THEN subst], rule foldSetD.intros)
ballarin@13936
    81
    apply auto
ballarin@13936
    82
  done
ballarin@13936
    83
ballarin@13936
    84
lemma (in LCD) foldSetD_imp_finite [simp]:
paulson@14750
    85
  "(A, x) \<in> foldSetD D f e ==> finite A"
ballarin@13936
    86
  by (induct set: foldSetD) auto
ballarin@13936
    87
ballarin@13936
    88
lemma (in LCD) finite_imp_foldSetD:
paulson@14750
    89
  "[| finite A; A \<subseteq> B; e \<in> D |] ==> EX x. (A, x) \<in> foldSetD D f e"
ballarin@13936
    90
proof (induct set: Finites)
ballarin@13936
    91
  case empty then show ?case by auto
ballarin@13936
    92
next
nipkow@15328
    93
  case (insert x F)
paulson@14750
    94
  then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto
paulson@14750
    95
  with insert have "y \<in> D" by auto
paulson@14750
    96
  with y and insert have "(insert x F, f x y) \<in> foldSetD D f e"
ballarin@13936
    97
    by (intro foldSetD.intros) auto
ballarin@13936
    98
  then show ?case ..
ballarin@13936
    99
qed
ballarin@13936
   100
ballarin@13936
   101
lemma (in LCD) foldSetD_determ_aux:
paulson@14750
   102
  "e \<in> D ==> \<forall>A x. A \<subseteq> B & card A < n --> (A, x) \<in> foldSetD D f e -->
paulson@14750
   103
    (\<forall>y. (A, y) \<in> foldSetD D f e --> y = x)"
ballarin@13936
   104
  apply (induct n)
ballarin@13936
   105
   apply (auto simp add: less_Suc_eq) (* slow *)
ballarin@13936
   106
  apply (erule foldSetD.cases)
ballarin@13936
   107
   apply blast
ballarin@13936
   108
  apply (erule foldSetD.cases)
ballarin@13936
   109
   apply blast
ballarin@13936
   110
  apply clarify
ballarin@13936
   111
  txt {* force simplification of @{text "card A < card (insert ...)"}. *}
ballarin@13936
   112
  apply (erule rev_mp)
ballarin@13936
   113
  apply (simp add: less_Suc_eq_le)
ballarin@13936
   114
  apply (rule impI)
ballarin@13936
   115
  apply (rename_tac Aa xa ya Ab xb yb, case_tac "xa = xb")
ballarin@13936
   116
   apply (subgoal_tac "Aa = Ab")
ballarin@13936
   117
    prefer 2 apply (blast elim!: equalityE)
ballarin@13936
   118
   apply blast
ballarin@13936
   119
  txt {* case @{prop "xa \<notin> xb"}. *}
paulson@14750
   120
  apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb \<in> Aa & xa \<in> Ab")
ballarin@13936
   121
   prefer 2 apply (blast elim!: equalityE)
ballarin@13936
   122
  apply clarify
ballarin@13936
   123
  apply (subgoal_tac "Aa = insert xb Ab - {xa}")
ballarin@13936
   124
   prefer 2 apply blast
paulson@14750
   125
  apply (subgoal_tac "card Aa \<le> card Ab")
ballarin@13936
   126
   prefer 2
ballarin@13936
   127
   apply (rule Suc_le_mono [THEN subst])
ballarin@13936
   128
   apply (simp add: card_Suc_Diff1)
ballarin@13936
   129
  apply (rule_tac A1 = "Aa - {xb}" in finite_imp_foldSetD [THEN exE])
ballarin@13936
   130
     apply (blast intro: foldSetD_imp_finite finite_Diff)
ballarin@13936
   131
    apply best
ballarin@13936
   132
   apply assumption
ballarin@13936
   133
  apply (frule (1) Diff1_foldSetD)
ballarin@13936
   134
   apply best
ballarin@13936
   135
  apply (subgoal_tac "ya = f xb x")
ballarin@13936
   136
   prefer 2
paulson@14750
   137
   apply (subgoal_tac "Aa \<subseteq> B")
ballarin@13936
   138
    prefer 2 apply best (* slow *)
ballarin@13936
   139
   apply (blast del: equalityCE)
paulson@14750
   140
  apply (subgoal_tac "(Ab - {xa}, x) \<in> foldSetD D f e")
ballarin@13936
   141
   prefer 2 apply simp
ballarin@13936
   142
  apply (subgoal_tac "yb = f xa x")
ballarin@13936
   143
   prefer 2 
ballarin@13936
   144
   apply (blast del: equalityCE dest: Diff1_foldSetD)
ballarin@13936
   145
  apply (simp (no_asm_simp))
ballarin@13936
   146
  apply (rule left_commute)
ballarin@13936
   147
    apply assumption
ballarin@13936
   148
   apply best (* slow *)
ballarin@13936
   149
  apply best
ballarin@13936
   150
  done
ballarin@13936
   151
ballarin@13936
   152
lemma (in LCD) foldSetD_determ:
paulson@14750
   153
  "[| (A, x) \<in> foldSetD D f e; (A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B |]
ballarin@13936
   154
  ==> y = x"
ballarin@13936
   155
  by (blast intro: foldSetD_determ_aux [rule_format])
ballarin@13936
   156
ballarin@13936
   157
lemma (in LCD) foldD_equality:
paulson@14750
   158
  "[| (A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B |] ==> foldD D f e A = y"
ballarin@13936
   159
  by (unfold foldD_def) (blast intro: foldSetD_determ)
ballarin@13936
   160
ballarin@13936
   161
lemma foldD_empty [simp]:
paulson@14750
   162
  "e \<in> D ==> foldD D f e {} = e"
ballarin@13936
   163
  by (unfold foldD_def) blast
ballarin@13936
   164
ballarin@13936
   165
lemma (in LCD) foldD_insert_aux:
paulson@14750
   166
  "[| x ~: A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
paulson@14750
   167
    ((insert x A, v) \<in> foldSetD D f e) =
paulson@14750
   168
    (EX y. (A, y) \<in> foldSetD D f e & v = f x y)"
ballarin@13936
   169
  apply auto
ballarin@13936
   170
  apply (rule_tac A1 = A in finite_imp_foldSetD [THEN exE])
ballarin@13936
   171
     apply (fastsimp dest: foldSetD_imp_finite)
ballarin@13936
   172
    apply assumption
ballarin@13936
   173
   apply assumption
ballarin@13936
   174
  apply (blast intro: foldSetD_determ)
ballarin@13936
   175
  done
ballarin@13936
   176
ballarin@13936
   177
lemma (in LCD) foldD_insert:
paulson@14750
   178
    "[| finite A; x ~: A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
ballarin@13936
   179
     foldD D f e (insert x A) = f x (foldD D f e A)"
ballarin@13936
   180
  apply (unfold foldD_def)
ballarin@13936
   181
  apply (simp add: foldD_insert_aux)
ballarin@13936
   182
  apply (rule the_equality)
ballarin@13936
   183
   apply (auto intro: finite_imp_foldSetD
ballarin@13936
   184
     cong add: conj_cong simp add: foldD_def [symmetric] foldD_equality)
ballarin@13936
   185
  done
ballarin@13936
   186
ballarin@13936
   187
lemma (in LCD) foldD_closed [simp]:
paulson@14750
   188
  "[| finite A; e \<in> D; A \<subseteq> B |] ==> foldD D f e A \<in> D"
ballarin@13936
   189
proof (induct set: Finites)
ballarin@13936
   190
  case empty then show ?case by (simp add: foldD_empty)
ballarin@13936
   191
next
ballarin@13936
   192
  case insert then show ?case by (simp add: foldD_insert)
ballarin@13936
   193
qed
ballarin@13936
   194
ballarin@13936
   195
lemma (in LCD) foldD_commute:
paulson@14750
   196
  "[| finite A; x \<in> B; e \<in> D; A \<subseteq> B |] ==>
ballarin@13936
   197
   f x (foldD D f e A) = foldD D f (f x e) A"
ballarin@13936
   198
  apply (induct set: Finites)
ballarin@13936
   199
   apply simp
ballarin@13936
   200
  apply (auto simp add: left_commute foldD_insert)
ballarin@13936
   201
  done
ballarin@13936
   202
ballarin@13936
   203
lemma Int_mono2:
paulson@14750
   204
  "[| A \<subseteq> C; B \<subseteq> C |] ==> A Int B \<subseteq> C"
ballarin@13936
   205
  by blast
ballarin@13936
   206
ballarin@13936
   207
lemma (in LCD) foldD_nest_Un_Int:
paulson@14750
   208
  "[| finite A; finite C; e \<in> D; A \<subseteq> B; C \<subseteq> B |] ==>
ballarin@13936
   209
   foldD D f (foldD D f e C) A = foldD D f (foldD D f e (A Int C)) (A Un C)"
ballarin@13936
   210
  apply (induct set: Finites)
ballarin@13936
   211
   apply simp
ballarin@13936
   212
  apply (simp add: foldD_insert foldD_commute Int_insert_left insert_absorb
ballarin@13936
   213
    Int_mono2 Un_subset_iff)
ballarin@13936
   214
  done
ballarin@13936
   215
ballarin@13936
   216
lemma (in LCD) foldD_nest_Un_disjoint:
paulson@14750
   217
  "[| finite A; finite B; A Int B = {}; e \<in> D; A \<subseteq> B; C \<subseteq> B |]
ballarin@13936
   218
    ==> foldD D f e (A Un B) = foldD D f (foldD D f e B) A"
ballarin@13936
   219
  by (simp add: foldD_nest_Un_Int)
ballarin@13936
   220
ballarin@13936
   221
-- {* Delete rules to do with @{text foldSetD} relation. *}
ballarin@13936
   222
ballarin@13936
   223
declare foldSetD_imp_finite [simp del]
ballarin@13936
   224
  empty_foldSetDE [rule del]
ballarin@13936
   225
  foldSetD.intros [rule del]
ballarin@13936
   226
declare (in LCD)
ballarin@13936
   227
  foldSetD_closed [rule del]
ballarin@13936
   228
ballarin@13936
   229
subsection {* Commutative monoids *}
ballarin@13936
   230
ballarin@13936
   231
text {*
ballarin@13936
   232
  We enter a more restrictive context, with @{text "f :: 'a => 'a => 'a"}
ballarin@13936
   233
  instead of @{text "'b => 'a => 'a"}.
ballarin@13936
   234
*}
ballarin@13936
   235
ballarin@13936
   236
locale ACeD =
ballarin@13936
   237
  fixes D :: "'a set"
ballarin@13936
   238
    and f :: "'a => 'a => 'a"    (infixl "\<cdot>" 70)
ballarin@13936
   239
    and e :: 'a
paulson@14750
   240
  assumes ident [simp]: "x \<in> D ==> x \<cdot> e = x"
paulson@14750
   241
    and commute: "[| x \<in> D; y \<in> D |] ==> x \<cdot> y = y \<cdot> x"
paulson@14750
   242
    and assoc: "[| x \<in> D; y \<in> D; z \<in> D |] ==> (x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)"
paulson@14750
   243
    and e_closed [simp]: "e \<in> D"
paulson@14750
   244
    and f_closed [simp]: "[| x \<in> D; y \<in> D |] ==> x \<cdot> y \<in> D"
ballarin@13936
   245
ballarin@13936
   246
lemma (in ACeD) left_commute:
paulson@14750
   247
  "[| x \<in> D; y \<in> D; z \<in> D |] ==> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
ballarin@13936
   248
proof -
paulson@14750
   249
  assume D: "x \<in> D" "y \<in> D" "z \<in> D"
ballarin@13936
   250
  then have "x \<cdot> (y \<cdot> z) = (y \<cdot> z) \<cdot> x" by (simp add: commute)
ballarin@13936
   251
  also from D have "... = y \<cdot> (z \<cdot> x)" by (simp add: assoc)
ballarin@13936
   252
  also from D have "z \<cdot> x = x \<cdot> z" by (simp add: commute)
ballarin@13936
   253
  finally show ?thesis .
ballarin@13936
   254
qed
ballarin@13936
   255
ballarin@13936
   256
lemmas (in ACeD) AC = assoc commute left_commute
ballarin@13936
   257
paulson@14750
   258
lemma (in ACeD) left_ident [simp]: "x \<in> D ==> e \<cdot> x = x"
ballarin@13936
   259
proof -
paulson@14750
   260
  assume D: "x \<in> D"
ballarin@13936
   261
  have "x \<cdot> e = x" by (rule ident)
ballarin@13936
   262
  with D show ?thesis by (simp add: commute)
ballarin@13936
   263
qed
ballarin@13936
   264
ballarin@13936
   265
lemma (in ACeD) foldD_Un_Int:
paulson@14750
   266
  "[| finite A; finite B; A \<subseteq> D; B \<subseteq> D |] ==>
ballarin@13936
   267
    foldD D f e A \<cdot> foldD D f e B =
ballarin@13936
   268
    foldD D f e (A Un B) \<cdot> foldD D f e (A Int B)"
ballarin@13936
   269
  apply (induct set: Finites)
ballarin@13936
   270
   apply (simp add: left_commute LCD.foldD_closed [OF LCD.intro [of D]])
ballarin@13936
   271
  apply (simp add: AC insert_absorb Int_insert_left
ballarin@13936
   272
    LCD.foldD_insert [OF LCD.intro [of D]]
ballarin@13936
   273
    LCD.foldD_closed [OF LCD.intro [of D]]
ballarin@13936
   274
    Int_mono2 Un_subset_iff)
ballarin@13936
   275
  done
ballarin@13936
   276
ballarin@13936
   277
lemma (in ACeD) foldD_Un_disjoint:
paulson@14750
   278
  "[| finite A; finite B; A Int B = {}; A \<subseteq> D; B \<subseteq> D |] ==>
ballarin@13936
   279
    foldD D f e (A Un B) = foldD D f e A \<cdot> foldD D f e B"
ballarin@13936
   280
  by (simp add: foldD_Un_Int
ballarin@13936
   281
    left_commute LCD.foldD_closed [OF LCD.intro [of D]] Un_subset_iff)
ballarin@13936
   282
ballarin@13936
   283
subsection {* Products over Finite Sets *}
ballarin@13936
   284
wenzelm@14651
   285
constdefs (structure G)
ballarin@15095
   286
  finprod :: "[('b, 'm) monoid_scheme, 'a => 'b, 'a set] => 'b"
ballarin@13936
   287
  "finprod G f A == if finite A
wenzelm@14651
   288
      then foldD (carrier G) (mult G o f) \<one> A
ballarin@13936
   289
      else arbitrary"
ballarin@13936
   290
wenzelm@14651
   291
syntax
wenzelm@14651
   292
  "_finprod" :: "index => idt => 'a set => 'b => 'b"
wenzelm@14666
   293
      ("(3\<Otimes>__:_. _)" [1000, 0, 51, 10] 10)
wenzelm@14651
   294
syntax (xsymbols)
wenzelm@14651
   295
  "_finprod" :: "index => idt => 'a set => 'b => 'b"
wenzelm@14666
   296
      ("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10)
wenzelm@14651
   297
syntax (HTML output)
wenzelm@14651
   298
  "_finprod" :: "index => idt => 'a set => 'b => 'b"
wenzelm@14666
   299
      ("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10)
wenzelm@14651
   300
translations
ballarin@15095
   301
  "\<Otimes>\<index>i:A. b" == "finprod \<struct>\<index> (%i. b) A"
ballarin@15095
   302
  -- {* Beware of argument permutation! *}
ballarin@13936
   303
ballarin@13936
   304
ML_setup {* 
wenzelm@14590
   305
  simpset_ref() := simpset() setsubgoaler asm_full_simp_tac;
ballarin@13936
   306
*}
ballarin@13936
   307
ballarin@13936
   308
lemma (in comm_monoid) finprod_empty [simp]: 
ballarin@13936
   309
  "finprod G f {} = \<one>"
ballarin@13936
   310
  by (simp add: finprod_def)
ballarin@13936
   311
ballarin@13936
   312
ML_setup {* 
wenzelm@14590
   313
  simpset_ref() := simpset() setsubgoaler asm_simp_tac;
ballarin@13936
   314
*}
ballarin@13936
   315
ballarin@13936
   316
declare funcsetI [intro]
ballarin@13936
   317
  funcset_mem [dest]
ballarin@13936
   318
ballarin@13936
   319
lemma (in comm_monoid) finprod_insert [simp]:
ballarin@13936
   320
  "[| finite F; a \<notin> F; f \<in> F -> carrier G; f a \<in> carrier G |] ==>
ballarin@13936
   321
   finprod G f (insert a F) = f a \<otimes> finprod G f F"
ballarin@13936
   322
  apply (rule trans)
ballarin@13936
   323
   apply (simp add: finprod_def)
ballarin@13936
   324
  apply (rule trans)
ballarin@13936
   325
   apply (rule LCD.foldD_insert [OF LCD.intro [of "insert a F"]])
ballarin@13936
   326
         apply simp
ballarin@13936
   327
         apply (rule m_lcomm)
ballarin@13936
   328
           apply fast
ballarin@13936
   329
          apply fast
ballarin@13936
   330
         apply assumption
ballarin@13936
   331
        apply (fastsimp intro: m_closed)
ballarin@13936
   332
       apply simp+
ballarin@13936
   333
   apply fast
ballarin@13936
   334
  apply (auto simp add: finprod_def)
ballarin@13936
   335
  done
ballarin@13936
   336
ballarin@13936
   337
lemma (in comm_monoid) finprod_one [simp]:
wenzelm@14651
   338
  "finite A ==> (\<Otimes>i:A. \<one>) = \<one>"
ballarin@13936
   339
proof (induct set: Finites)
ballarin@13936
   340
  case empty show ?case by simp
ballarin@13936
   341
next
nipkow@15328
   342
  case (insert a A)
ballarin@13936
   343
  have "(%i. \<one>) \<in> A -> carrier G" by auto
ballarin@13936
   344
  with insert show ?case by simp
ballarin@13936
   345
qed
ballarin@13936
   346
ballarin@13936
   347
lemma (in comm_monoid) finprod_closed [simp]:
ballarin@13936
   348
  fixes A
ballarin@13936
   349
  assumes fin: "finite A" and f: "f \<in> A -> carrier G" 
ballarin@13936
   350
  shows "finprod G f A \<in> carrier G"
ballarin@13936
   351
using fin f
ballarin@13936
   352
proof induct
ballarin@13936
   353
  case empty show ?case by simp
ballarin@13936
   354
next
nipkow@15328
   355
  case (insert a A)
ballarin@13936
   356
  then have a: "f a \<in> carrier G" by fast
ballarin@13936
   357
  from insert have A: "f \<in> A -> carrier G" by fast
ballarin@13936
   358
  from insert A a show ?case by simp
ballarin@13936
   359
qed
ballarin@13936
   360
ballarin@13936
   361
lemma funcset_Int_left [simp, intro]:
ballarin@13936
   362
  "[| f \<in> A -> C; f \<in> B -> C |] ==> f \<in> A Int B -> C"
ballarin@13936
   363
  by fast
ballarin@13936
   364
ballarin@13936
   365
lemma funcset_Un_left [iff]:
ballarin@13936
   366
  "(f \<in> A Un B -> C) = (f \<in> A -> C & f \<in> B -> C)"
ballarin@13936
   367
  by fast
ballarin@13936
   368
ballarin@13936
   369
lemma (in comm_monoid) finprod_Un_Int:
ballarin@13936
   370
  "[| finite A; finite B; g \<in> A -> carrier G; g \<in> B -> carrier G |] ==>
ballarin@13936
   371
     finprod G g (A Un B) \<otimes> finprod G g (A Int B) =
ballarin@13936
   372
     finprod G g A \<otimes> finprod G g B"
ballarin@13936
   373
-- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
ballarin@13936
   374
proof (induct set: Finites)
ballarin@13936
   375
  case empty then show ?case by (simp add: finprod_closed)
ballarin@13936
   376
next
nipkow@15328
   377
  case (insert a A)
ballarin@13936
   378
  then have a: "g a \<in> carrier G" by fast
ballarin@13936
   379
  from insert have A: "g \<in> A -> carrier G" by fast
ballarin@13936
   380
  from insert A a show ?case
ballarin@13936
   381
    by (simp add: m_ac Int_insert_left insert_absorb finprod_closed
ballarin@13936
   382
          Int_mono2 Un_subset_iff) 
ballarin@13936
   383
qed
ballarin@13936
   384
ballarin@13936
   385
lemma (in comm_monoid) finprod_Un_disjoint:
ballarin@13936
   386
  "[| finite A; finite B; A Int B = {};
ballarin@13936
   387
      g \<in> A -> carrier G; g \<in> B -> carrier G |]
ballarin@13936
   388
   ==> finprod G g (A Un B) = finprod G g A \<otimes> finprod G g B"
ballarin@13936
   389
  apply (subst finprod_Un_Int [symmetric])
ballarin@13936
   390
      apply (auto simp add: finprod_closed)
ballarin@13936
   391
  done
ballarin@13936
   392
ballarin@13936
   393
lemma (in comm_monoid) finprod_multf:
ballarin@13936
   394
  "[| finite A; f \<in> A -> carrier G; g \<in> A -> carrier G |] ==>
ballarin@13936
   395
   finprod G (%x. f x \<otimes> g x) A = (finprod G f A \<otimes> finprod G g A)"
ballarin@13936
   396
proof (induct set: Finites)
ballarin@13936
   397
  case empty show ?case by simp
ballarin@13936
   398
next
nipkow@15328
   399
  case (insert a A) then
paulson@14750
   400
  have fA: "f \<in> A -> carrier G" by fast
paulson@14750
   401
  from insert have fa: "f a \<in> carrier G" by fast
paulson@14750
   402
  from insert have gA: "g \<in> A -> carrier G" by fast
paulson@14750
   403
  from insert have ga: "g a \<in> carrier G" by fast
paulson@14750
   404
  from insert have fgA: "(%x. f x \<otimes> g x) \<in> A -> carrier G"
ballarin@13936
   405
    by (simp add: Pi_def)
ballarin@15095
   406
  show ?case
ballarin@15095
   407
    by (simp add: insert fA fa gA ga fgA m_ac)
ballarin@13936
   408
qed
ballarin@13936
   409
ballarin@13936
   410
lemma (in comm_monoid) finprod_cong':
paulson@14750
   411
  "[| A = B; g \<in> B -> carrier G;
paulson@14750
   412
      !!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B"
ballarin@13936
   413
proof -
paulson@14750
   414
  assume prems: "A = B" "g \<in> B -> carrier G"
paulson@14750
   415
    "!!i. i \<in> B ==> f i = g i"
ballarin@13936
   416
  show ?thesis
ballarin@13936
   417
  proof (cases "finite B")
ballarin@13936
   418
    case True
paulson@14750
   419
    then have "!!A. [| A = B; g \<in> B -> carrier G;
paulson@14750
   420
      !!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B"
ballarin@13936
   421
    proof induct
ballarin@13936
   422
      case empty thus ?case by simp
ballarin@13936
   423
    next
nipkow@15328
   424
      case (insert x B)
ballarin@13936
   425
      then have "finprod G f A = finprod G f (insert x B)" by simp
ballarin@13936
   426
      also from insert have "... = f x \<otimes> finprod G f B"
ballarin@13936
   427
      proof (intro finprod_insert)
ballarin@13936
   428
	show "finite B" .
ballarin@13936
   429
      next
ballarin@13936
   430
	show "x ~: B" .
ballarin@13936
   431
      next
ballarin@13936
   432
	assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
ballarin@13936
   433
	  "g \<in> insert x B \<rightarrow> carrier G"
paulson@14750
   434
	thus "f \<in> B -> carrier G" by fastsimp
ballarin@13936
   435
      next
ballarin@13936
   436
	assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
ballarin@13936
   437
	  "g \<in> insert x B \<rightarrow> carrier G"
ballarin@13936
   438
	thus "f x \<in> carrier G" by fastsimp
ballarin@13936
   439
      qed
ballarin@13936
   440
      also from insert have "... = g x \<otimes> finprod G g B" by fastsimp
ballarin@13936
   441
      also from insert have "... = finprod G g (insert x B)"
ballarin@13936
   442
      by (intro finprod_insert [THEN sym]) auto
ballarin@13936
   443
      finally show ?case .
ballarin@13936
   444
    qed
ballarin@13936
   445
    with prems show ?thesis by simp
ballarin@13936
   446
  next
ballarin@13936
   447
    case False with prems show ?thesis by (simp add: finprod_def)
ballarin@13936
   448
  qed
ballarin@13936
   449
qed
ballarin@13936
   450
ballarin@13936
   451
lemma (in comm_monoid) finprod_cong:
paulson@14750
   452
  "[| A = B; f \<in> B -> carrier G = True;
paulson@14750
   453
      !!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B"
ballarin@14213
   454
  (* This order of prems is slightly faster (3%) than the last two swapped. *)
ballarin@14213
   455
  by (rule finprod_cong') force+
ballarin@13936
   456
ballarin@13936
   457
text {*Usually, if this rule causes a failed congruence proof error,
paulson@14750
   458
  the reason is that the premise @{text "g \<in> B -> carrier G"} cannot be shown.
ballarin@13936
   459
  Adding @{thm [source] Pi_def} to the simpset is often useful.
ballarin@13936
   460
  For this reason, @{thm [source] comm_monoid.finprod_cong}
ballarin@13936
   461
  is not added to the simpset by default.
ballarin@13936
   462
*}
ballarin@13936
   463
ballarin@13936
   464
declare funcsetI [rule del]
ballarin@13936
   465
  funcset_mem [rule del]
ballarin@13936
   466
ballarin@13936
   467
lemma (in comm_monoid) finprod_0 [simp]:
paulson@14750
   468
  "f \<in> {0::nat} -> carrier G ==> finprod G f {..0} = f 0"
ballarin@13936
   469
by (simp add: Pi_def)
ballarin@13936
   470
ballarin@13936
   471
lemma (in comm_monoid) finprod_Suc [simp]:
paulson@14750
   472
  "f \<in> {..Suc n} -> carrier G ==>
ballarin@13936
   473
   finprod G f {..Suc n} = (f (Suc n) \<otimes> finprod G f {..n})"
ballarin@13936
   474
by (simp add: Pi_def atMost_Suc)
ballarin@13936
   475
ballarin@13936
   476
lemma (in comm_monoid) finprod_Suc2:
paulson@14750
   477
  "f \<in> {..Suc n} -> carrier G ==>
ballarin@13936
   478
   finprod G f {..Suc n} = (finprod G (%i. f (Suc i)) {..n} \<otimes> f 0)"
ballarin@13936
   479
proof (induct n)
ballarin@13936
   480
  case 0 thus ?case by (simp add: Pi_def)
ballarin@13936
   481
next
ballarin@13936
   482
  case Suc thus ?case by (simp add: m_assoc Pi_def)
ballarin@13936
   483
qed
ballarin@13936
   484
ballarin@13936
   485
lemma (in comm_monoid) finprod_mult [simp]:
paulson@14750
   486
  "[| f \<in> {..n} -> carrier G; g \<in> {..n} -> carrier G |] ==>
ballarin@13936
   487
     finprod G (%i. f i \<otimes> g i) {..n::nat} =
ballarin@13936
   488
     finprod G f {..n} \<otimes> finprod G g {..n}"
ballarin@13936
   489
  by (induct n) (simp_all add: m_ac Pi_def)
ballarin@13936
   490
ballarin@13936
   491
end
ballarin@13936
   492