haftmann@26168: (* Title: HOL/Library/RType.thy haftmann@26168: ID: $Id$ haftmann@26168: Author: Florian Haftmann, TU Muenchen haftmann@26168: *) haftmann@26168: haftmann@26168: header {* Reflecting Pure types into HOL *} haftmann@26168: haftmann@28952: theory Typerep haftmann@28952: imports Plain List Code_Message haftmann@26168: begin haftmann@26168: haftmann@28335: datatype typerep = Typerep message_string "typerep list" haftmann@26168: haftmann@28335: class typerep = haftmann@28335: fixes typerep :: "'a\{} itself \ typerep" haftmann@26168: begin haftmann@26168: haftmann@26168: definition haftmann@28335: typerep_of :: "'a \ typerep" haftmann@26168: where haftmann@28335: [simp]: "typerep_of x = typerep TYPE('a)" haftmann@26168: haftmann@26168: end haftmann@26168: haftmann@26168: setup {* haftmann@26168: let haftmann@28335: fun typerep_tr (*"_TYPEREP"*) [ty] = haftmann@28335: Lexicon.const @{const_syntax typerep} $ (Lexicon.const "_constrain" $ Lexicon.const "TYPE" $ haftmann@26168: (Lexicon.const "itself" $ ty)) haftmann@28335: | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts); haftmann@28335: fun typerep_tr' show_sorts (*"typerep"*) haftmann@26168: (Type ("fun", [Type ("itself", [T]), _])) (Const (@{const_syntax TYPE}, _) :: ts) = haftmann@28335: Term.list_comb (Lexicon.const "_TYPEREP" $ Syntax.term_of_typ show_sorts T, ts) haftmann@28335: | typerep_tr' _ T ts = raise Match; haftmann@26168: in haftmann@26168: Sign.add_syntax_i haftmann@28335: [("_TYPEREP", SimpleSyntax.read_typ "type => logic", Delimfix "(1TYPEREP/(1'(_')))")] haftmann@28335: #> Sign.add_trfuns ([], [("_TYPEREP", typerep_tr)], [], []) haftmann@28335: #> Sign.add_trfunsT [(@{const_syntax typerep}, typerep_tr')] haftmann@26168: end haftmann@26168: *} haftmann@26168: haftmann@26168: ML {* haftmann@28335: structure Typerep = haftmann@26168: struct haftmann@26168: haftmann@26168: fun mk f (Type (tyco, tys)) = haftmann@28335: @{term Typerep} $ Message_String.mk tyco haftmann@28335: $ HOLogic.mk_list @{typ typerep} (map (mk f) tys) haftmann@26168: | mk f (TFree v) = haftmann@26168: f v; haftmann@26168: haftmann@28335: fun typerep ty = haftmann@28335: Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep}) haftmann@26168: $ Logic.mk_type ty; haftmann@26168: haftmann@26168: fun add_def tyco thy = haftmann@26168: let haftmann@28335: val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep}; haftmann@26168: val vs = Name.names Name.context "'a" sorts; haftmann@26168: val ty = Type (tyco, map TFree vs); haftmann@28335: val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep}) haftmann@26168: $ Free ("T", Term.itselfT ty); haftmann@28335: val rhs = mk (typerep o TFree) ty; haftmann@26168: val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); haftmann@26168: in haftmann@26168: thy haftmann@28335: |> TheoryTarget.instantiation ([tyco], vs, @{sort typerep}) haftmann@26168: |> `(fn lthy => Syntax.check_term lthy eq) haftmann@28965: |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq))) haftmann@26168: |> snd haftmann@26168: |> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) wenzelm@28394: |> LocalTheory.exit_global haftmann@26168: end; haftmann@26168: haftmann@26168: fun perhaps_add_def tyco thy = haftmann@26168: let haftmann@28335: val inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep} haftmann@26168: in if inst then thy else add_def tyco thy end; haftmann@26168: haftmann@26168: end; haftmann@26168: *} haftmann@26168: haftmann@26168: setup {* haftmann@28335: Typerep.add_def @{type_name prop} haftmann@28335: #> Typerep.add_def @{type_name fun} haftmann@28335: #> Typerep.add_def @{type_name itself} haftmann@28335: #> Typerep.add_def @{type_name bool} haftmann@28335: #> TypedefPackage.interpretation Typerep.perhaps_add_def haftmann@26168: *} haftmann@26168: haftmann@28562: lemma [code]: haftmann@28346: "eq_class.eq (Typerep tyco1 tys1) (Typerep tyco2 tys2) \ eq_class.eq tyco1 tyco2 haftmann@28346: \ list_all2 eq_class.eq tys1 tys2" haftmann@28346: by (auto simp add: equals_eq [symmetric] list_all2_eq [symmetric]) haftmann@26168: haftmann@28335: code_type typerep haftmann@26168: (SML "Term.typ") haftmann@26168: haftmann@28335: code_const Typerep haftmann@26168: (SML "Term.Type/ (_, _)") haftmann@26168: haftmann@26168: code_reserved SML Term haftmann@26168: haftmann@28335: hide (open) const typerep Typerep haftmann@26168: haftmann@26168: end