Folds.thy
author urbanc
Thu, 11 Aug 2011 23:42:06 +0000
changeset 194 5347d7556487
parent 85 a3e0056c228b
child 203 5d724fe0e096
permissions -rw-r--r--
some typos
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
85
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     1
theory Folds
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     2
imports Main
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     3
begin
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     4
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     5
section {* Folds for Sets *}
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     6
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     7
text {*
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     8
  To obtain equational system out of finite set of equivalence classes, a fold operation
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
     9
  on finite sets @{text "folds"} is defined. The use of @{text "SOME"} makes @{text "folds"}
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    10
  more robust than the @{text "fold"} in the Isabelle library. The expression @{text "folds f"}
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    11
  makes sense when @{text "f"} is not @{text "associative"} and @{text "commutitive"},
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    12
  while @{text "fold f"} does not.  
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    13
*}
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    14
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    15
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    16
definition 
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    17
  folds :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    18
where
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    19
  "folds f z S \<equiv> SOME x. fold_graph f z S x"
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    20
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    21
a3e0056c228b separated the definition of folds into a separate file
urbanc
parents:
diff changeset
    22
end