src/HOL/Record.thy
author schirmer
Mon, 03 May 2004 23:22:17 +0200
changeset 14700 2f885b7e5ba7
parent 14080 9a50427d7165
child 14701 62a724ce51c7
permissions -rw-r--r--
reimplementation of HOL records; only one type is created for each record extension, instead of one type for each field. See NEWS.
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$
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
     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
11826
2203c7f9ec40 proper setup for abstract product types;
wenzelm
parents: 11821
diff changeset
     6
theory Record = Product_Type
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
     7
files ("Tools/record_package.ML"):
4870
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
     8
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
     9
ML {*
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    10
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
    11
by (rtac h2 1);
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    12
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
    13
qed "meta_allE";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    14
*}
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    15
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    16
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
    17
  by simp
11826
2203c7f9ec40 proper setup for abstract product types;
wenzelm
parents: 11821
diff changeset
    18
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    19
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
    20
  by simp
11826
2203c7f9ec40 proper setup for abstract product types;
wenzelm
parents: 11821
diff changeset
    21
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    22
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
    23
  by simp
14080
9a50427d7165 Added split_paired_All rule for splitting variables bound by
berghofe
parents: 13421
diff changeset
    24
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    25
11833
wenzelm
parents: 11826
diff changeset
    26
subsection {* Concrete record syntax *}
4870
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
    27
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
    28
nonterminals
5198
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    29
  ident field_type field_types field fields update updates
4870
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
    30
syntax
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    31
  "_constify"           :: "id => ident"                        ("_")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    32
  "_constify"           :: "longid => ident"                    ("_")
5198
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    33
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    34
  "_field_type"         :: "[ident, type] => field_type"        ("(2_ ::/ _)")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    35
  ""                    :: "field_type => field_types"          ("_")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    36
  "_field_types"        :: "[field_type, field_types] => field_types"    ("_,/ _")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    37
  "_record_type"        :: "field_types => type"                ("(3'(| _ |'))")
10093
44584c2b512b more symbolic syntax (currently "input");
wenzelm
parents: 9729
diff changeset
    38
  "_record_type_scheme" :: "[field_types, type] => type"        ("(3'(| _,/ (2... ::/ _) |'))")
5198
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    39
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    40
  "_field"              :: "[ident, 'a] => field"               ("(2_ =/ _)")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    41
  ""                    :: "field => fields"                    ("_")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    42
  "_fields"             :: "[field, fields] => fields"          ("_,/ _")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    43
  "_record"             :: "fields => 'a"                       ("(3'(| _ |'))")
10093
44584c2b512b more symbolic syntax (currently "input");
wenzelm
parents: 9729
diff changeset
    44
  "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3'(| _,/ (2... =/ _) |'))")
5198
b1adae4f8b90 added type and update syntax;
wenzelm
parents: 5032
diff changeset
    45
10641
d1533f63c738 added "_update_name" and parse_translation;
wenzelm
parents: 10331
diff changeset
    46
  "_update_name"        :: idt
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    47
  "_update"             :: "[ident, 'a] => update"              ("(2_ :=/ _)")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    48
  ""                    :: "update => updates"                  ("_")
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    49
  "_updates"            :: "[update, updates] => updates"       ("_,/ _")
10093
44584c2b512b more symbolic syntax (currently "input");
wenzelm
parents: 9729
diff changeset
    50
  "_record_update"      :: "['a, updates] => 'b"                ("_/(3'(| _ |'))" [900,0] 900)
4870
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
    51
10331
7411e4659d4a more "xsymbols" syntax;
wenzelm
parents: 10309
diff changeset
    52
syntax (xsymbols)
11821
ad32c92435db abstract product types;
wenzelm
parents: 11489
diff changeset
    53
  "_record_type"        :: "field_types => type"                ("(3\<lparr>_\<rparr>)")
10093
44584c2b512b more symbolic syntax (currently "input");
wenzelm
parents: 9729
diff changeset
    54
  "_record_type_scheme" :: "[field_types, type] => type"        ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)")
44584c2b512b more symbolic syntax (currently "input");
wenzelm
parents: 9729
diff changeset
    55
  "_record"             :: "fields => 'a"                               ("(3\<lparr>_\<rparr>)")
44584c2b512b more symbolic syntax (currently "input");
wenzelm
parents: 9729
diff changeset
    56
  "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
44584c2b512b more symbolic syntax (currently "input");
wenzelm
parents: 9729
diff changeset
    57
  "_record_update"      :: "['a, updates] => 'b"                ("_/(3\<lparr>_\<rparr>)" [900,0] 900)
9729
40cfc3dd27da \<dots> syntax;
wenzelm
parents: 7357
diff changeset
    58
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    59
(* 
11833
wenzelm
parents: 11826
diff changeset
    60
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    61
  "_structure"             :: "fields => 'a"          ("(3{| _ |})")
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    62
  "_structure_scheme"      :: "[fields, 'a] => 'a"    ("(3{| _,/ (2... =/ _) |})")
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    63
  
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    64
  "_structure_update_name":: idt
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    65
  "_structure_update"  :: "['a, updates] \<Rightarrow> 'b"    ("_/(3{| _ |})" [900,0] 900)
11833
wenzelm
parents: 11826
diff changeset
    66
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    67
  "_structure_type"        :: "field_types => type"    ("(3{| _ |})")
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    68
  "_structure_type_scheme" :: "[field_types, type] => type" 
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    69
                                      ("(3{| _,/ (2... ::/ _) |})")
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    70
syntax (xsymbols)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    71
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    72
 "_structure_scheme"   :: "[fields, 'a] => 'a"       ("(3{|_,/ (2\<dots> =/ _)|})")
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    73
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    74
  "_structure_type_scheme" :: "[field_types, type] => type"        
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    75
                                      ("(3{|_,/ (2\<dots> ::/ _)|})")
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    76
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    77
*)
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    78
use "Tools/record_package.ML";
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    79
setup RecordPackage.setup;
10641
d1533f63c738 added "_update_name" and parse_translation;
wenzelm
parents: 10331
diff changeset
    80
4870
cc36acb5b114 Extensible records with structural subtyping in HOL. See
wenzelm
parents:
diff changeset
    81
end
14700
2f885b7e5ba7 reimplementation of HOL records; only one type is created for
schirmer
parents: 14080
diff changeset
    82