| author | wenzelm | 
| Mon, 29 Feb 2016 22:34:36 +0100 | |
| changeset 62479 | 716336f19aa9 | 
| parent 62175 | 8ffc4d0e652d | 
| child 65379 | 76a96e32bd23 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Plain_HOLCF.thy | 
| 40502 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 2 | Author: Brian Huffman | 
| 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 3 | *) | 
| 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 4 | |
| 62175 | 5 | section \<open>Plain HOLCF\<close> | 
| 40502 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 6 | |
| 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 7 | theory Plain_HOLCF | 
| 40592 
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
 huffman parents: 
40502diff
changeset | 8 | imports Cfun Sfun Cprod Sprod Ssum Up Discrete Lift One Tr Fix | 
| 40502 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 9 | begin | 
| 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 10 | |
| 62175 | 11 | text \<open> | 
| 40502 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 12 | Basic HOLCF concepts and types; does not include definition packages. | 
| 62175 | 13 | \<close> | 
| 40502 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 14 | |
| 60040 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
58880diff
changeset | 15 | hide_const (open) Filter.principal | 
| 
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
 hoelzl parents: 
58880diff
changeset | 16 | |
| 40502 
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
 huffman parents: diff
changeset | 17 | end |