src/ZF/Induct/FoldSet.thy
author huffman
Tue, 02 Mar 2010 09:54:50 -0800
changeset 35512 d1ef88d7de5a
parent 32960 69916a850301
child 46822 95f1e700b712
permissions -rw-r--r--
remove dead code
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12610
8b9845807f77 tuned document sources;
wenzelm
parents: 12089
diff changeset
     1
(*  Title:      ZF/Induct/FoldSet.thy
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     2
    Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     3
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     4
A "fold" functional for finite sets.  For n non-negative we have
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     5
fold f e {x1,...,xn} = f x1 (... (f xn e)) where f is at
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     6
least left-commutative.  
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     7
*)
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14071
diff changeset
     9
theory FoldSet imports Main begin
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    10
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    11
consts fold_set :: "[i, i, [i,i]=>i, i] => i"
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    12
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    13
inductive
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    14
  domains "fold_set(A, B, f,e)" <= "Fin(A)*B"
14071
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    15
  intros
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    16
    emptyI: "e\<in>B ==> <0, e>\<in>fold_set(A, B, f,e)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    17
    consI:  "[| x\<in>A; x \<notin>C;  <C,y> : fold_set(A, B,f,e); f(x,y):B |]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
    18
                ==>  <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)"
14071
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    19
  type_intros Fin.intros
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    20
  
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    21
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    22
  fold :: "[i, [i,i]=>i, i, i] => i"  ("fold[_]'(_,_,_')")  where
14071
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    23
   "fold[B](f,e, A) == THE x. <A, x>\<in>fold_set(A, B, f,e)"
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
    24
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    25
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    26
   setsum :: "[i=>i, i] => i"  where
14071
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    27
   "setsum(g, C) == if Finite(C) then
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    28
                     fold[int](%x y. g(x) $+ y, #0, C) else #0"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    29
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    30
(** foldSet **)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    31
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    32
inductive_cases empty_fold_setE: "<0, x> : fold_set(A, B, f,e)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    33
inductive_cases cons_fold_setE: "<cons(x,C), y> : fold_set(A, B, f,e)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    34
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    35
(* add-hoc lemmas *)                                                
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    36
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    37
lemma cons_lemma1: "[| x\<notin>C; x\<notin>B |] ==> cons(x,B)=cons(x,C) <-> B = C"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    38
by (auto elim: equalityE)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    39
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    40
lemma cons_lemma2: "[| cons(x, B)=cons(y, C); x\<noteq>y; x\<notin>B; y\<notin>C |]  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    41
    ==>  B - {y} = C-{x} & x\<in>C & y\<in>B"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    42
apply (auto elim: equalityE)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    43
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    44
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    45
(* fold_set monotonicity *)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    46
lemma fold_set_mono_lemma:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    47
     "<C, x> : fold_set(A, B, f, e)  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    48
      ==> ALL D. A<=D --> <C, x> : fold_set(D, B, f, e)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    49
apply (erule fold_set.induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    50
apply (auto intro: fold_set.intros)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    51
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    52
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    53
lemma fold_set_mono: " C<=A ==> fold_set(C, B, f, e) <= fold_set(A, B, f, e)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    54
apply clarify
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    55
apply (frule fold_set.dom_subset [THEN subsetD], clarify)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    56
apply (auto dest: fold_set_mono_lemma)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    57
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    58
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    59
lemma fold_set_lemma:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    60
     "<C, x>\<in>fold_set(A, B, f, e) ==> <C, x>\<in>fold_set(C, B, f, e) & C<=A"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    61
apply (erule fold_set.induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    62
apply (auto intro!: fold_set.intros intro: fold_set_mono [THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    63
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    64
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    65
(* Proving that fold_set is deterministic *)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    66
lemma Diff1_fold_set:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    67
     "[| <C-{x},y> : fold_set(A, B, f,e);  x\<in>C; x\<in>A; f(x, y):B |]  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    68
      ==> <C, f(x, y)> : fold_set(A, B, f, e)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    69
apply (frule fold_set.dom_subset [THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    70
apply (erule cons_Diff [THEN subst], rule fold_set.intros, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    71
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    72
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    73
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    74
locale fold_typing =
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    75
 fixes A and B and e and f
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    76
 assumes ftype [intro,simp]:  "[|x \<in> A; y \<in> B|] ==> f(x,y) \<in> B"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    77
     and etype [intro,simp]:  "e \<in> B"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    78
     and fcomm:  "[|x \<in> A; y \<in> A; z \<in> B|] ==> f(x, f(y, z))=f(y, f(x, z))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    79
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    80
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    81
lemma (in fold_typing) Fin_imp_fold_set:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    82
     "C\<in>Fin(A) ==> (EX x. <C, x> : fold_set(A, B, f,e))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    83
apply (erule Fin_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    84
apply (auto dest: fold_set.dom_subset [THEN subsetD] 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    85
            intro: fold_set.intros etype ftype)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    86
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    87
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    88
lemma Diff_sing_imp:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    89
     "[|C - {b} = D - {a}; a \<noteq> b; b \<in> C|] ==> C = cons(b,D) - {a}"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    90
by (blast elim: equalityE)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    91
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    92
lemma (in fold_typing) fold_set_determ_lemma [rule_format]: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    93
"n\<in>nat
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    94
 ==> ALL C. |C|<n -->  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    95
   (ALL x. <C, x> : fold_set(A, B, f,e)--> 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    96
           (ALL y. <C, y> : fold_set(A, B, f,e) --> y=x))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    97
apply (erule nat_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    98
 apply (auto simp add: le_iff)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
    99
apply (erule fold_set.cases)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   100
 apply (force elim!: empty_fold_setE)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   101
apply (erule fold_set.cases)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   102
 apply (force elim!: empty_fold_setE, clarify)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   103
(*force simplification of "|C| < |cons(...)|"*)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   104
apply (frule_tac a = Ca in fold_set.dom_subset [THEN subsetD, THEN SigmaD1])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   105
apply (frule_tac a = Cb in fold_set.dom_subset [THEN subsetD, THEN SigmaD1])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   106
apply (simp add: Fin_into_Finite [THEN Finite_imp_cardinal_cons])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   107
apply (case_tac "x=xb", auto) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   108
apply (simp add: cons_lemma1, blast)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   109
txt{*case @{term "x\<noteq>xb"}*}
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   110
apply (drule cons_lemma2, safe)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   111
apply (frule Diff_sing_imp, assumption+) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   112
txt{** LEVEL 17*}
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   113
apply (subgoal_tac "|Ca| le |Cb|")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   114
 prefer 2
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   115
 apply (rule succ_le_imp_le)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   116
 apply (simp add: Fin_into_Finite Finite_imp_succ_cardinal_Diff 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   117
                  Fin_into_Finite [THEN Finite_imp_cardinal_cons])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   118
apply (rule_tac C1 = "Ca-{xb}" in Fin_imp_fold_set [THEN exE])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   119
 apply (blast intro: Diff_subset [THEN Fin_subset])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   120
txt{** LEVEL 24 **}
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   121
apply (frule Diff1_fold_set, blast, blast)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   122
apply (blast dest!: ftype fold_set.dom_subset [THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   123
apply (subgoal_tac "ya = f(xb,xa) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   124
 prefer 2 apply (blast del: equalityCE)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   125
apply (subgoal_tac "<Cb-{x}, xa> : fold_set(A,B,f,e)")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   126
 prefer 2 apply simp
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   127
apply (subgoal_tac "yb = f (x, xa) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   128
 apply (drule_tac [2] C = Cb in Diff1_fold_set, simp_all)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   129
  apply (blast intro: fcomm dest!: fold_set.dom_subset [THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   130
 apply (blast intro: ftype dest!: fold_set.dom_subset [THEN subsetD], blast) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   131
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   132
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   133
lemma (in fold_typing) fold_set_determ: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   134
     "[| <C, x>\<in>fold_set(A, B, f, e);  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   135
         <C, y>\<in>fold_set(A, B, f, e)|] ==> y=x"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   136
apply (frule fold_set.dom_subset [THEN subsetD], clarify)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   137
apply (drule Fin_into_Finite)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   138
apply (unfold Finite_def, clarify)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   139
apply (rule_tac n = "succ (n)" in fold_set_determ_lemma) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   140
apply (auto intro: eqpoll_imp_lepoll [THEN lepoll_cardinal_le])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   141
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   142
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   143
(** The fold function **)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   144
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   145
lemma (in fold_typing) fold_equality: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   146
     "<C,y> : fold_set(A,B,f,e) ==> fold[B](f,e,C) = y"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   147
apply (unfold fold_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   148
apply (frule fold_set.dom_subset [THEN subsetD], clarify)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   149
apply (rule the_equality)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   150
 apply (rule_tac [2] A=C in fold_typing.fold_set_determ)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   151
apply (force dest: fold_set_lemma)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   152
apply (auto dest: fold_set_lemma)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   153
apply (simp add: fold_typing_def, auto) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   154
apply (auto dest: fold_set_lemma intro: ftype etype fcomm)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   155
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   156
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   157
lemma fold_0 [simp]: "e : B ==> fold[B](f,e,0) = e"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   158
apply (unfold fold_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   159
apply (blast elim!: empty_fold_setE intro: fold_set.intros)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   160
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   161
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   162
text{*This result is the right-to-left direction of the subsequent result*}
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   163
lemma (in fold_typing) fold_set_imp_cons: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   164
     "[| <C, y> : fold_set(C, B, f, e); C : Fin(A); c : A; c\<notin>C |]
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   165
      ==> <cons(c, C), f(c,y)> : fold_set(cons(c, C), B, f, e)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   166
apply (frule FinD [THEN fold_set_mono, THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   167
 apply assumption
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   168
apply (frule fold_set.dom_subset [of A, THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   169
apply (blast intro!: fold_set.consI intro: fold_set_mono [THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   170
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   171
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   172
lemma (in fold_typing) fold_cons_lemma [rule_format]: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   173
"[| C : Fin(A); c : A; c\<notin>C |]   
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   174
     ==> <cons(c, C), v> : fold_set(cons(c, C), B, f, e) <->   
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   175
         (EX y. <C, y> : fold_set(C, B, f, e) & v = f(c, y))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   176
apply auto
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   177
 prefer 2 apply (blast intro: fold_set_imp_cons) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   178
 apply (frule_tac Fin.consI [of c, THEN FinD, THEN fold_set_mono, THEN subsetD], assumption+)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   179
apply (frule_tac fold_set.dom_subset [of A, THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   180
apply (drule FinD) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   181
apply (rule_tac A1 = "cons(c,C)" and f1=f and B1=B and C1=C and e1=e in fold_typing.Fin_imp_fold_set [THEN exE])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   182
apply (blast intro: fold_typing.intro ftype etype fcomm) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   183
apply (blast intro: Fin_subset [of _ "cons(c,C)"] Finite_into_Fin 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   184
             dest: Fin_into_Finite)  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   185
apply (rule_tac x = x in exI)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   186
apply (auto intro: fold_set.intros)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   187
apply (drule_tac fold_set_lemma [of C], blast)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   188
apply (blast intro!: fold_set.consI
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   189
             intro: fold_set_determ fold_set_mono [THEN subsetD] 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   190
             dest: fold_set.dom_subset [THEN subsetD]) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   191
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   192
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   193
lemma (in fold_typing) fold_cons: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   194
     "[| C\<in>Fin(A); c\<in>A; c\<notin>C|] 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   195
      ==> fold[B](f, e, cons(c, C)) = f(c, fold[B](f, e, C))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   196
apply (unfold fold_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   197
apply (simp add: fold_cons_lemma)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   198
apply (rule the_equality, auto) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   199
 apply (subgoal_tac [2] "\<langle>C, y\<rangle> \<in> fold_set(A, B, f, e)")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   200
  apply (drule Fin_imp_fold_set)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   201
apply (auto dest: fold_set_lemma  simp add: fold_def [symmetric] fold_equality) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   202
apply (blast intro: fold_set_mono [THEN subsetD] dest!: FinD) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   203
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   204
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   205
lemma (in fold_typing) fold_type [simp,TC]: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   206
     "C\<in>Fin(A) ==> fold[B](f,e,C):B"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   207
apply (erule Fin_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   208
apply (simp_all add: fold_cons ftype etype)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   209
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   210
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   211
lemma (in fold_typing) fold_commute [rule_format]: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   212
     "[| C\<in>Fin(A); c\<in>A |]  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   213
      ==> (\<forall>y\<in>B. f(c, fold[B](f, y, C)) = fold[B](f, f(c, y), C))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   214
apply (erule Fin_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   215
apply (simp_all add: fold_typing.fold_cons [of A B _ f] 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   216
                     fold_typing.fold_type [of A B _ f] 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   217
                     fold_typing_def fcomm)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   218
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   219
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   220
lemma (in fold_typing) fold_nest_Un_Int: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   221
     "[| C\<in>Fin(A); D\<in>Fin(A) |]
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   222
      ==> fold[B](f, fold[B](f, e, D), C) =
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   223
          fold[B](f, fold[B](f, e, (C Int D)), C Un D)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   224
apply (erule Fin_induct, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   225
apply (simp add: Un_cons Int_cons_left fold_type fold_commute
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   226
                 fold_typing.fold_cons [of A _ _ f] 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   227
                 fold_typing_def fcomm cons_absorb)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   228
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   229
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   230
lemma (in fold_typing) fold_nest_Un_disjoint:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   231
     "[| C\<in>Fin(A); D\<in>Fin(A); C Int D = 0 |]  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   232
      ==> fold[B](f,e,C Un D) =  fold[B](f, fold[B](f,e,D), C)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   233
by (simp add: fold_nest_Un_Int)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   234
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   235
lemma Finite_cons_lemma: "Finite(C) ==> C\<in>Fin(cons(c, C))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   236
apply (drule Finite_into_Fin)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   237
apply (blast intro: Fin_mono [THEN subsetD])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   238
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   239
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   240
subsection{*The Operator @{term setsum}*}
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   241
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   242
lemma setsum_0 [simp]: "setsum(g, 0) = #0"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   243
by (simp add: setsum_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   244
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   245
lemma setsum_cons [simp]: 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   246
     "Finite(C) ==> 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   247
      setsum(g, cons(c,C)) = 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   248
        (if c : C then setsum(g,C) else g(c) $+ setsum(g,C))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   249
apply (auto simp add: setsum_def Finite_cons cons_absorb)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   250
apply (rule_tac A = "cons (c, C)" in fold_typing.fold_cons)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   251
apply (auto intro: fold_typing.intro Finite_cons_lemma)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   252
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   253
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   254
lemma setsum_K0: "setsum((%i. #0), C) = #0"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   255
apply (case_tac "Finite (C) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   256
 prefer 2 apply (simp add: setsum_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   257
apply (erule Finite_induct, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   258
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   259
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   260
(*The reversed orientation looks more natural, but LOOPS as a simprule!*)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   261
lemma setsum_Un_Int:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   262
     "[| Finite(C); Finite(D) |]  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   263
      ==> setsum(g, C Un D) $+ setsum(g, C Int D)  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   264
        = setsum(g, C) $+ setsum(g, D)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   265
apply (erule Finite_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   266
apply (simp_all add: Int_cons_right cons_absorb Un_cons Int_commute Finite_Un
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   267
                     Int_lower1 [THEN subset_Finite]) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   268
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   269
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   270
lemma setsum_type [simp,TC]: "setsum(g, C):int"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   271
apply (case_tac "Finite (C) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   272
 prefer 2 apply (simp add: setsum_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   273
apply (erule Finite_induct, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   274
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   275
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   276
lemma setsum_Un_disjoint:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   277
     "[| Finite(C); Finite(D); C Int D = 0 |]  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   278
      ==> setsum(g, C Un D) = setsum(g, C) $+ setsum(g,D)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   279
apply (subst setsum_Un_Int [symmetric])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   280
apply (subgoal_tac [3] "Finite (C Un D) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   281
apply (auto intro: Finite_Un)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   282
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   283
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   284
lemma Finite_RepFun [rule_format (no_asm)]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   285
     "Finite(I) ==> (\<forall>i\<in>I. Finite(C(i))) --> Finite(RepFun(I, C))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   286
apply (erule Finite_induct, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   287
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   288
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   289
lemma setsum_UN_disjoint [rule_format (no_asm)]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   290
     "Finite(I)  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   291
      ==> (\<forall>i\<in>I. Finite(C(i))) -->  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   292
          (\<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j --> C(i) Int C(j) = 0) -->  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   293
          setsum(f, \<Union>i\<in>I. C(i)) = setsum (%i. setsum(f, C(i)), I)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   294
apply (erule Finite_induct, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   295
apply (subgoal_tac "\<forall>i\<in>B. x \<noteq> i")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   296
 prefer 2 apply blast
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   297
apply (subgoal_tac "C (x) Int (\<Union>i\<in>B. C (i)) = 0")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   298
 prefer 2 apply blast
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   299
apply (subgoal_tac "Finite (\<Union>i\<in>B. C (i)) & Finite (C (x)) & Finite (B) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   300
apply (simp (no_asm_simp) add: setsum_Un_disjoint)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   301
apply (auto intro: Finite_Union Finite_RepFun)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   302
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   303
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   304
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   305
lemma setsum_addf: "setsum(%x. f(x) $+ g(x),C) = setsum(f, C) $+ setsum(g, C)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   306
apply (case_tac "Finite (C) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   307
 prefer 2 apply (simp add: setsum_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   308
apply (erule Finite_induct, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   309
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   310
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   311
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   312
lemma fold_set_cong:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   313
     "[| A=A'; B=B'; e=e'; (\<forall>x\<in>A'. \<forall>y\<in>B'. f(x,y) = f'(x,y)) |] 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   314
      ==> fold_set(A,B,f,e) = fold_set(A',B',f',e')"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   315
apply (simp add: fold_set_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   316
apply (intro refl iff_refl lfp_cong Collect_cong disj_cong ex_cong, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   317
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   318
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   319
lemma fold_cong:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   320
"[| B=B'; A=A'; e=e';   
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   321
    !!x y. [|x\<in>A'; y\<in>B'|] ==> f(x,y) = f'(x,y) |] ==>  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   322
   fold[B](f,e,A) = fold[B'](f', e', A')"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   323
apply (simp add: fold_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   324
apply (subst fold_set_cong)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   325
apply (rule_tac [5] refl, simp_all)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   326
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   327
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   328
lemma setsum_cong:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   329
 "[| A=B; !!x. x\<in>B ==> f(x) = g(x) |] ==>  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   330
     setsum(f, A) = setsum(g, B)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   331
by (simp add: setsum_def cong add: fold_cong)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   332
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   333
lemma setsum_Un:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   334
     "[| Finite(A); Finite(B) |]  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   335
      ==> setsum(f, A Un B) =  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   336
          setsum(f, A) $+ setsum(f, B) $- setsum(f, A Int B)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   337
apply (subst setsum_Un_Int [symmetric], auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   338
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   339
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   340
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   341
lemma setsum_zneg_or_0 [rule_format (no_asm)]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   342
     "Finite(A) ==> (\<forall>x\<in>A. g(x) $<= #0) --> setsum(g, A) $<= #0"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   343
apply (erule Finite_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   344
apply (auto intro: zneg_or_0_add_zneg_or_0_imp_zneg_or_0)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   345
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   346
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   347
lemma setsum_succD_lemma [rule_format]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   348
     "Finite(A)  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   349
      ==> \<forall>n\<in>nat. setsum(f,A) = $# succ(n) --> (\<exists>a\<in>A. #0 $< f(a))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   350
apply (erule Finite_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   351
apply (auto simp del: int_of_0 int_of_succ simp add: not_zless_iff_zle int_of_0 [symmetric])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   352
apply (subgoal_tac "setsum (f, B) $<= #0")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   353
apply simp_all
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   354
prefer 2 apply (blast intro: setsum_zneg_or_0)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   355
apply (subgoal_tac "$# 1 $<= f (x) $+ setsum (f, B) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   356
apply (drule zdiff_zle_iff [THEN iffD2])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   357
apply (subgoal_tac "$# 1 $<= $# 1 $- setsum (f,B) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   358
apply (drule_tac x = "$# 1" in zle_trans)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   359
apply (rule_tac [2] j = "#1" in zless_zle_trans, auto)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   360
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   361
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   362
lemma setsum_succD:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   363
     "[| setsum(f, A) = $# succ(n); n\<in>nat |]==> \<exists>a\<in>A. #0 $< f(a)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   364
apply (case_tac "Finite (A) ")
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   365
apply (blast intro: setsum_succD_lemma)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   366
apply (unfold setsum_def)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   367
apply (auto simp del: int_of_0 int_of_succ simp add: int_succ_int_1 [symmetric] int_of_0 [symmetric])
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   368
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   369
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   370
lemma g_zpos_imp_setsum_zpos [rule_format]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   371
     "Finite(A) ==> (\<forall>x\<in>A. #0 $<= g(x)) --> #0 $<= setsum(g, A)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   372
apply (erule Finite_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   373
apply (simp (no_asm))
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   374
apply (auto intro: zpos_add_zpos_imp_zpos)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   375
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   376
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   377
lemma g_zpos_imp_setsum_zpos2 [rule_format]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   378
     "[| Finite(A); \<forall>x. #0 $<= g(x) |] ==> #0 $<= setsum(g, A)"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   379
apply (erule Finite_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   380
apply (auto intro: zpos_add_zpos_imp_zpos)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   381
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   382
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   383
lemma g_zspos_imp_setsum_zspos [rule_format]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   384
     "Finite(A)  
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   385
      ==> (\<forall>x\<in>A. #0 $< g(x)) --> A \<noteq> 0 --> (#0 $< setsum(g, A))"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   386
apply (erule Finite_induct)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   387
apply (auto intro: zspos_add_zspos_imp_zspos)
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   388
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   389
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   390
lemma setsum_Diff [rule_format]:
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   391
     "Finite(A) ==> \<forall>a. M(a) = #0 --> setsum(M, A) = setsum(M, A-{a})"
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   392
apply (erule Finite_induct) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   393
apply (simp_all add: Diff_cons_eq Finite_Diff) 
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   394
done
373806545656 Converting ZF/UNITY to Isar
paulson
parents: 14046
diff changeset
   395
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
diff changeset
   396
end