| author | wenzelm | 
| Fri, 11 Oct 2013 20:45:21 +0200 | |
| changeset 54325 | 2c4155003352 | 
| parent 54026 | 82d9b2701a03 | 
| child 54539 | bbab2ebda234 | 
| permissions | -rw-r--r-- | 
| 
49509
 
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
 
blanchet 
parents: 
49309 
diff
changeset
 | 
1  | 
(* Title: HOL/BNF/BNF.thy  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
| 49017 | 3  | 
Author: Andrei Popescu, TU Muenchen  | 
4  | 
Author: Jasmin Blanchette, TU Muenchen  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
Copyright 2012  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
|
| 
49509
 
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
 
blanchet 
parents: 
49309 
diff
changeset
 | 
7  | 
Bounded natural functors for (co)datatypes.  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
9  | 
|
| 
49509
 
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
 
blanchet 
parents: 
49309 
diff
changeset
 | 
10  | 
header {* Bounded Natural Functors for (Co)datatypes *}
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
11  | 
|
| 
49509
 
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
 
blanchet 
parents: 
49309 
diff
changeset
 | 
12  | 
theory BNF  | 
| 54026 | 13  | 
imports More_BNFs BNF_LFP BNF_GFP Coinduction  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
14  | 
begin  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
15  | 
|
| 52731 | 16  | 
hide_const (open) image2 image2p vimage2p Gr Grp collect fsts snds setl setr  | 
| 51909 | 17  | 
convol thePull pick_middlep fstOp sndOp csquare inver  | 
| 51446 | 18  | 
image2 relImage relInvImage prefCl PrefCl Succ Shift shift  | 
19  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
20  | 
end  |