| author | wenzelm | 
| Mon, 30 May 2005 23:07:58 +0200 | |
| changeset 16114 | 8d453f906e43 | 
| parent 15140 | 322485b816ac | 
| child 16417 | 9bc16273c2d4 | 
| 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  | 
| 15131 | 8  | 
files ("Tools/record_package.ML")
 | 
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: 
14080 
diff
changeset
 | 
11  | 
ML {*
 | 
| 
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
12  | 
val [h1, h2] = Goal "PROP Goal (\<And>x. PROP P x) \<Longrightarrow> (PROP P x \<Longrightarrow> PROP Q) \<Longrightarrow> PROP Q";  | 
| 
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
13  | 
by (rtac h2 1);  | 
| 
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
14  | 
by (rtac (gen_all (h1 RS Drule.rev_triv_goal)) 1);  | 
| 
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
15  | 
qed "meta_allE";  | 
| 
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
16  | 
*}  | 
| 11821 | 17  | 
|
| 
14700
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
18  | 
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: 
14080 
diff
changeset
 | 
19  | 
by simp  | 
| 11826 | 20  | 
|
| 
14700
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
21  | 
lemma rec_UNIV_I: "\<And>x. x\<in>UNIV \<equiv> True"  | 
| 
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
22  | 
by simp  | 
| 11826 | 23  | 
|
| 
14700
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
24  | 
lemma rec_True_simp: "(True \<Longrightarrow> PROP P) \<equiv> PROP P"  | 
| 
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
25  | 
by simp  | 
| 
14080
 
9a50427d7165
Added split_paired_All rule for splitting variables bound by
 
berghofe 
parents: 
13421 
diff
changeset
 | 
26  | 
|
| 11821 | 27  | 
|
| 11833 | 28  | 
subsection {* Concrete record syntax *}
 | 
| 
4870
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
nonterminals  | 
| 5198 | 31  | 
ident field_type field_types field fields update updates  | 
| 
4870
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
syntax  | 
| 11821 | 33  | 
  "_constify"           :: "id => ident"                        ("_")
 | 
34  | 
  "_constify"           :: "longid => ident"                    ("_")
 | 
|
| 5198 | 35  | 
|
| 11821 | 36  | 
  "_field_type"         :: "[ident, type] => field_type"        ("(2_ ::/ _)")
 | 
37  | 
  ""                    :: "field_type => field_types"          ("_")
 | 
|
38  | 
  "_field_types"        :: "[field_type, field_types] => field_types"    ("_,/ _")
 | 
|
39  | 
  "_record_type"        :: "field_types => type"                ("(3'(| _ |'))")
 | 
|
| 10093 | 40  | 
  "_record_type_scheme" :: "[field_types, type] => type"        ("(3'(| _,/ (2... ::/ _) |'))")
 | 
| 5198 | 41  | 
|
| 11821 | 42  | 
  "_field"              :: "[ident, 'a] => field"               ("(2_ =/ _)")
 | 
43  | 
  ""                    :: "field => fields"                    ("_")
 | 
|
44  | 
  "_fields"             :: "[field, fields] => fields"          ("_,/ _")
 | 
|
45  | 
  "_record"             :: "fields => 'a"                       ("(3'(| _ |'))")
 | 
|
| 10093 | 46  | 
  "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3'(| _,/ (2... =/ _) |'))")
 | 
| 5198 | 47  | 
|
| 10641 | 48  | 
"_update_name" :: idt  | 
| 11821 | 49  | 
  "_update"             :: "[ident, 'a] => update"              ("(2_ :=/ _)")
 | 
50  | 
  ""                    :: "update => updates"                  ("_")
 | 
|
51  | 
  "_updates"            :: "[update, updates] => updates"       ("_,/ _")
 | 
|
| 10093 | 52  | 
  "_record_update"      :: "['a, updates] => 'b"                ("_/(3'(| _ |'))" [900,0] 900)
 | 
| 
4870
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
|
| 10331 | 54  | 
syntax (xsymbols)  | 
| 11821 | 55  | 
  "_record_type"        :: "field_types => type"                ("(3\<lparr>_\<rparr>)")
 | 
| 10093 | 56  | 
  "_record_type_scheme" :: "[field_types, type] => type"        ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)")
 | 
57  | 
  "_record"             :: "fields => 'a"                               ("(3\<lparr>_\<rparr>)")
 | 
|
58  | 
  "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
 | 
|
59  | 
  "_record_update"      :: "['a, updates] => 'b"                ("_/(3\<lparr>_\<rparr>)" [900,0] 900)
 | 
|
| 9729 | 60  | 
|
| 16114 | 61  | 
use "Tools/record_package.ML"  | 
62  | 
setup RecordPackage.setup  | 
|
| 10641 | 63  | 
|
| 
4870
 
cc36acb5b114
Extensible records with structural subtyping in HOL.  See
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
end  | 
| 
14700
 
2f885b7e5ba7
reimplementation of HOL records; only one type is created for
 
schirmer 
parents: 
14080 
diff
changeset
 | 
65  |