| 20400 |      1 | (*  ID:         $Id$
 | 
|  |      2 |     Author:     Florian Haftmann, TU Muenchen
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | header {* Combinators for structured results *}
 | 
|  |      6 | 
 | 
|  |      7 | theory CodeRevappl
 | 
|  |      8 | imports Main
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | section {* Combinators for structured results *}
 | 
|  |     12 | 
 | 
| 20707 |     13 | 
 | 
|  |     14 | subsection {* primitive definitions *}
 | 
|  |     15 | 
 | 
| 20400 |     16 | definition
 | 
|  |     17 |   revappl :: "'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b" (infixl "\<triangleright>" 55)
 | 
|  |     18 |   revappl_def: "x \<triangleright> f = f x"
 | 
|  |     19 |   revappl_snd :: "'c \<times> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'c \<times> 'b" (infixl "|\<triangleright>" 55)
 | 
|  |     20 |   revappl_snd_split: "z |\<triangleright> f = (fst z , f (snd z))"
 | 
|  |     21 |   revappl_swamp :: "'c \<times> 'a \<Rightarrow> ('a \<Rightarrow> 'd \<times> 'b) \<Rightarrow> ('c \<times> 'd) \<times> 'b" (infixl "|\<triangleright>\<triangleright>" 55)
 | 
|  |     22 |   revappl_swamp_split: "z |\<triangleright>\<triangleright> f = ((fst z, fst (f (snd z))), snd (f (snd z)))"
 | 
|  |     23 |   revappl_uncurry :: "'c \<times> 'a \<Rightarrow> ('c \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'b" (infixl "\<turnstile>\<triangleright>" 55)
 | 
|  |     24 |   revappl_uncurry_split: "z \<turnstile>\<triangleright> f = f (fst z) (snd z)"
 | 
|  |     25 | 
 | 
| 20707 |     26 | 
 | 
|  |     27 | subsection {* lemmas *}
 | 
|  |     28 | 
 | 
| 20400 |     29 | lemma revappl_snd_def [code]:
 | 
|  |     30 |   "(y, x) |\<triangleright> f = (y, f x)" unfolding revappl_snd_split by simp
 | 
|  |     31 | 
 | 
|  |     32 | lemma revappl_swamp_def [code]:
 | 
|  |     33 |   "(y, x) |\<triangleright>\<triangleright> f = (let (z, w) = f x in ((y, z), w))" unfolding Let_def revappl_swamp_split split_def by simp
 | 
|  |     34 | 
 | 
|  |     35 | lemma revappl_uncurry_def [code]:
 | 
|  |     36 |   "(y, x) \<turnstile>\<triangleright> f = f y x" unfolding revappl_uncurry_split by simp
 | 
|  |     37 | 
 | 
|  |     38 | lemmas revappl = revappl_def revappl_snd_def revappl_swamp_def revappl_uncurry_def
 | 
|  |     39 | 
 | 
|  |     40 | lemmas revappl_split = revappl_def revappl_snd_split revappl_swamp_split revappl_uncurry_split
 | 
|  |     41 | 
 | 
|  |     42 | end |