| author | wenzelm | 
| Mon, 26 Jun 2000 00:00:40 +0200 | |
| changeset 9143 | 6180c29d2db6 | 
| parent 7357 | d0e16da40ea2 | 
| child 9729 | 40cfc3dd27da | 
| 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  | 
|
| 7357 | 9  | 
theory Record = Datatype  | 
10  | 
files "Tools/record_package.ML":  | 
|
| 
4870
 
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  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
(* concrete syntax *)  | 
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
nonterminals  | 
| 5198 | 16  | 
ident field_type field_types field fields update updates  | 
| 
4870
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
syntax  | 
| 5198 | 19  | 
(*field names*)  | 
20  | 
  ""                	:: "id => ident"                    		("_")
 | 
|
21  | 
  ""                	:: "longid => ident"                		("_")
 | 
|
22  | 
||
23  | 
(*record types*)  | 
|
24  | 
  "_field_type"         :: "[ident, type] => field_type"           	("(2_ ::/ _)")
 | 
|
25  | 
  ""               	:: "field_type => field_types"              	("_")
 | 
|
26  | 
  "_field_types"	:: "[field_type, field_types] => field_types"   ("_,/ _")
 | 
|
| 5732 | 27  | 
  "_record_type"	:: "field_types => type"                   	("(3'(| _ |'))")
 | 
28  | 
  "_record_type_scheme"	:: "[field_types, type] => type"	("(3'(| _,/ (2... ::/ _) |'))")
 | 
|
| 5198 | 29  | 
|
30  | 
(*records*)  | 
|
31  | 
  "_field"          	:: "[ident, 'a] => field"           		("(2_ =/ _)")
 | 
|
32  | 
  ""                	:: "field => fields"                		("_")
 | 
|
33  | 
  "_fields"         	:: "[field, fields] => fields"      		("_,/ _")
 | 
|
| 5732 | 34  | 
  "_record"         	:: "fields => 'a"                   		("(3'(| _ |'))")
 | 
35  | 
  "_record_scheme"  	:: "[fields, 'a] => 'a"             	("(3'(| _,/ (2... =/ _) |'))")
 | 
|
| 5198 | 36  | 
|
37  | 
(*record updates*)  | 
|
38  | 
  "_update"	    	:: "[ident, 'a] => update"           		("(2_ :=/ _)")
 | 
|
39  | 
  ""                	:: "update => updates"               		("_")
 | 
|
40  | 
  "_updates"        	:: "[update, updates] => updates"    		("_,/ _")
 | 
|
| 5732 | 41  | 
  "_record_update"  	:: "['a, updates] => 'b"               	("_/(3'(| _ |'))" [900,0] 900)
 | 
| 
4870
 
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  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
(* type class for record extensions *)  | 
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
|
| 7357 | 46  | 
axclass more < "term"  | 
| 
4870
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
instance unit :: more  | 
| 7357 | 48  | 
by intro_classes  | 
| 
4870
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
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  | 
(* initialize the package *)  | 
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
setup  | 
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
RecordPackage.setup  | 
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
end  |