Up to index of Isabelle/HOL
theory Record(* Title: HOL/Record.thy ID: $Id$ Author: Wolfgang Naraschewski, Norbert Schirmer and Markus Wenzel, TU Muenchen *) header {* Extensible records with structural subtyping *} theory Record imports Product_Type uses ("Tools/record_package.ML") begin lemma prop_subst: "s = t ==> PROP P t ==> PROP P s" by simp lemma rec_UNIV_I: "!!x. x∈UNIV ≡ True" by simp lemma rec_True_simp: "(True ==> PROP P) ≡ PROP P" by simp lemma K_record_comp: "(λx. c) o f = (λx. c)" by (simp add: comp_def) subsection {* Concrete record syntax *} nonterminals ident field_type field_types field fields update updates syntax "_constify" :: "id => ident" ("_") "_constify" :: "longid => ident" ("_") "_field_type" :: "[ident, type] => field_type" ("(2_ ::/ _)") "" :: "field_type => field_types" ("_") "_field_types" :: "[field_type, field_types] => field_types" ("_,/ _") "_record_type" :: "field_types => type" ("(3'(| _ |'))") "_record_type_scheme" :: "[field_types, type] => type" ("(3'(| _,/ (2... ::/ _) |'))") "_field" :: "[ident, 'a] => field" ("(2_ =/ _)") "" :: "field => fields" ("_") "_fields" :: "[field, fields] => fields" ("_,/ _") "_record" :: "fields => 'a" ("(3'(| _ |'))") "_record_scheme" :: "[fields, 'a] => 'a" ("(3'(| _,/ (2... =/ _) |'))") "_update_name" :: idt "_update" :: "[ident, 'a] => update" ("(2_ :=/ _)") "" :: "update => updates" ("_") "_updates" :: "[update, updates] => updates" ("_,/ _") "_record_update" :: "['a, updates] => 'b" ("_/(3'(| _ |'))" [900,0] 900) syntax (xsymbols) "_record_type" :: "field_types => type" ("(3(|_|)),)") "_record_type_scheme" :: "[field_types, type] => type" ("(3(|_,/ (2… ::/ _)|)),)") "_record" :: "fields => 'a" ("(3(|_|)),)") "_record_scheme" :: "[fields, 'a] => 'a" ("(3(|_,/ (2… =/ _)|)),)") "_record_update" :: "['a, updates] => 'b" ("_/(3(|_|)),)" [900,0] 900) use "Tools/record_package.ML" setup RecordPackage.setup end