| author | wenzelm | 
| Wed, 03 Feb 1999 16:45:45 +0100 | |
| changeset 6189 | e9dc9ec28a2d | 
| parent 5732 | 8712391bbf3d | 
| child 7357 | d0e16da40ea2 | 
| permissions | -rw-r--r-- | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 1 | (* Title: HOL/Record.thy | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 3 | Author: Wolfgang Naraschewski and Markus Wenzel, TU Muenchen | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 4 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 5 | Extensible records with structural subtyping in HOL. See | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 6 | Tools/record_package.ML for the actual implementation. | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 7 | *) | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 8 | |
| 5694 | 9 | Record = Datatype + | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 10 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 11 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 12 | (* concrete syntax *) | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 13 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 14 | nonterminals | 
| 5198 | 15 | ident field_type field_types field fields update updates | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 16 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 17 | syntax | 
| 5198 | 18 | (*field names*) | 
| 19 |   ""                	:: "id => ident"                    		("_")
 | |
| 20 |   ""                	:: "longid => ident"                		("_")
 | |
| 21 | ||
| 22 | (*record types*) | |
| 23 |   "_field_type"         :: "[ident, type] => field_type"           	("(2_ ::/ _)")
 | |
| 24 |   ""               	:: "field_type => field_types"              	("_")
 | |
| 25 |   "_field_types"	:: "[field_type, field_types] => field_types"   ("_,/ _")
 | |
| 5732 | 26 |   "_record_type"	:: "field_types => type"                   	("(3'(| _ |'))")
 | 
| 27 |   "_record_type_scheme"	:: "[field_types, type] => type"	("(3'(| _,/ (2... ::/ _) |'))")
 | |
| 5198 | 28 | |
| 29 | (*records*) | |
| 30 |   "_field"          	:: "[ident, 'a] => field"           		("(2_ =/ _)")
 | |
| 31 |   ""                	:: "field => fields"                		("_")
 | |
| 32 |   "_fields"         	:: "[field, fields] => fields"      		("_,/ _")
 | |
| 5732 | 33 |   "_record"         	:: "fields => 'a"                   		("(3'(| _ |'))")
 | 
| 34 |   "_record_scheme"  	:: "[fields, 'a] => 'a"             	("(3'(| _,/ (2... =/ _) |'))")
 | |
| 5198 | 35 | |
| 36 | (*record updates*) | |
| 37 |   "_update"	    	:: "[ident, 'a] => update"           		("(2_ :=/ _)")
 | |
| 38 |   ""                	:: "update => updates"               		("_")
 | |
| 39 |   "_updates"        	:: "[update, updates] => updates"    		("_,/ _")
 | |
| 5732 | 40 |   "_record_update"  	:: "['a, updates] => 'b"               	("_/(3'(| _ |'))" [900,0] 900)
 | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 41 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 42 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 43 | (* type class for record extensions *) | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 44 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 45 | axclass more < term | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 46 | instance unit :: more | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 47 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 48 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 49 | (* initialize the package *) | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 50 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 51 | setup | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 52 | RecordPackage.setup | 
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 53 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 54 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 55 | end |