src/HOL/Record.thy
author nipkow
Fri, 16 Oct 1998 17:36:12 +0200
changeset 5657 1a6c9c6a3f8e
parent 5210 54aaa779b6b4
child 5694 39af7b3dd1c4
permissions -rw-r--r--
2. The simplifier now knows a little bit about nat-arithmetic.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
     9
Record = Prod +
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
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    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
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    18
  (*field names*)
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    19
  ""                	:: "id => ident"                    		("_")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    20
  ""                	:: "longid => ident"                		("_")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    21
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    22
  (*record types*)
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    23
  "_field_type"         :: "[ident, type] => field_type"           	("(2_ ::/ _)")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    24
  ""               	:: "field_type => field_types"              	("_")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    25
  "_field_types"	:: "[field_type, field_types] => field_types"   ("_,/ _")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    26
  "_record_type"	:: "field_types => type"                   	("('(| _ |'))")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    27
  "_record_type_scheme"	:: "[field_types, type] => type"	("('(| _,/ (2... ::/ _) |'))")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    28
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    29
  (*records*)
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    30
  "_field"          	:: "[ident, 'a] => field"           		("(2_ =/ _)")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    31
  ""                	:: "field => fields"                		("_")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    32
  "_fields"         	:: "[field, fields] => fields"      		("_,/ _")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    33
  "_record"         	:: "fields => 'a"                   		("('(| _ |'))")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    34
  "_record_scheme"  	:: "[fields, 'a] => 'a"             	("('(| _,/ (2... =/ _) |'))")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    35
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    36
  (*record updates*)
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    37
  "_update"	    	:: "[ident, 'a] => update"           		("(2_ :=/ _)")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    38
  ""                	:: "update => updates"               		("_")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    39
  "_updates"        	:: "[update, updates] => updates"    		("_,/ _")
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    40
  "_record_update"  	:: "['a, updates] => 'b"               	("_/('(| _ |'))" [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
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
    47
instance unit :: more
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
    48
instance "*" :: (term, more) more
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