| author | berghofe | 
| Wed, 07 Feb 2007 18:00:38 +0100 | |
| changeset 22277 | b89dc456dbc6 | 
| parent 21226 | a607ae87ee81 | 
| child 22744 | 5cbe966d67a2 | 
| 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$ | 
| 16114 | 3 | Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen | 
| 4870 
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 | |
| 15131 | 6 | theory Record | 
| 15140 | 7 | imports Product_Type | 
| 16417 | 8 | uses ("Tools/record_package.ML")
 | 
| 15131 | 9 | begin | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 10 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 11 | lemma prop_subst: "s = t \<Longrightarrow> PROP P t \<Longrightarrow> PROP P s" | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 12 | by simp | 
| 11826 | 13 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 14 | lemma rec_UNIV_I: "\<And>x. x\<in>UNIV \<equiv> True" | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 15 | by simp | 
| 11826 | 16 | |
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 17 | lemma rec_True_simp: "(True \<Longrightarrow> PROP P) \<equiv> PROP P" | 
| 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 18 | by simp | 
| 14080 
9a50427d7165
Added split_paired_All rule for splitting variables bound by
 berghofe parents: 
13421diff
changeset | 19 | |
| 21226 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 20 | constdefs K_record:: "'a \<Rightarrow> 'b \<Rightarrow> 'a" | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 21 | "K_record c x \<equiv> c" | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 22 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 23 | lemma K_record_apply [simp]: "K_record c x = c" | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 24 | by (simp add: K_record_def) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 25 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 26 | lemma K_record_comp [simp]: "(K_record c \<circ> f) = K_record c" | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 27 | by (rule ext) (simp add: K_record_apply comp_def) | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 28 | |
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 29 | lemma K_record_cong [cong]: "K_record c x = K_record c x" | 
| 
a607ae87ee81
field-update in records is generalised to take a function on the field
 schirmer parents: 
17957diff
changeset | 30 | by (rule refl) | 
| 11821 | 31 | |
| 11833 | 32 | subsection {* Concrete record syntax *}
 | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 33 | |
| 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 34 | nonterminals | 
| 5198 | 35 | ident field_type field_types field fields update updates | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 36 | syntax | 
| 11821 | 37 |   "_constify"           :: "id => ident"                        ("_")
 | 
| 38 |   "_constify"           :: "longid => ident"                    ("_")
 | |
| 5198 | 39 | |
| 11821 | 40 |   "_field_type"         :: "[ident, type] => field_type"        ("(2_ ::/ _)")
 | 
| 41 |   ""                    :: "field_type => field_types"          ("_")
 | |
| 42 |   "_field_types"        :: "[field_type, field_types] => field_types"    ("_,/ _")
 | |
| 43 |   "_record_type"        :: "field_types => type"                ("(3'(| _ |'))")
 | |
| 10093 | 44 |   "_record_type_scheme" :: "[field_types, type] => type"        ("(3'(| _,/ (2... ::/ _) |'))")
 | 
| 5198 | 45 | |
| 11821 | 46 |   "_field"              :: "[ident, 'a] => field"               ("(2_ =/ _)")
 | 
| 47 |   ""                    :: "field => fields"                    ("_")
 | |
| 48 |   "_fields"             :: "[field, fields] => fields"          ("_,/ _")
 | |
| 49 |   "_record"             :: "fields => 'a"                       ("(3'(| _ |'))")
 | |
| 10093 | 50 |   "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3'(| _,/ (2... =/ _) |'))")
 | 
| 5198 | 51 | |
| 10641 | 52 | "_update_name" :: idt | 
| 11821 | 53 |   "_update"             :: "[ident, 'a] => update"              ("(2_ :=/ _)")
 | 
| 54 |   ""                    :: "update => updates"                  ("_")
 | |
| 55 |   "_updates"            :: "[update, updates] => updates"       ("_,/ _")
 | |
| 10093 | 56 |   "_record_update"      :: "['a, updates] => 'b"                ("_/(3'(| _ |'))" [900,0] 900)
 | 
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 57 | |
| 10331 | 58 | syntax (xsymbols) | 
| 11821 | 59 |   "_record_type"        :: "field_types => type"                ("(3\<lparr>_\<rparr>)")
 | 
| 10093 | 60 |   "_record_type_scheme" :: "[field_types, type] => type"        ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)")
 | 
| 61 |   "_record"             :: "fields => 'a"                               ("(3\<lparr>_\<rparr>)")
 | |
| 62 |   "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
 | |
| 63 |   "_record_update"      :: "['a, updates] => 'b"                ("_/(3\<lparr>_\<rparr>)" [900,0] 900)
 | |
| 9729 | 64 | |
| 16114 | 65 | use "Tools/record_package.ML" | 
| 66 | setup RecordPackage.setup | |
| 10641 | 67 | |
| 4870 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 wenzelm parents: diff
changeset | 68 | end | 
| 14700 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 schirmer parents: 
14080diff
changeset | 69 |