haftmann@31048: (* Author: Florian Haftmann, TU Muenchen *) haftmann@26168: haftmann@26168: header {* Reflecting Pure types into HOL *} haftmann@26168: haftmann@28952: theory Typerep haftmann@31048: imports Plain String haftmann@26168: begin haftmann@26168: haftmann@31205: datatype typerep = Typerep String.literal "typerep list" haftmann@26168: haftmann@28335: class typerep = haftmann@31031: fixes typerep :: "'a itself \ typerep" haftmann@26168: begin haftmann@26168: haftmann@29574: definition typerep_of :: "'a \ typerep" where haftmann@28335: [simp]: "typerep_of x = typerep TYPE('a)" haftmann@26168: haftmann@26168: end haftmann@26168: wenzelm@35115: syntax wenzelm@35115: "_TYPEREP" :: "type => logic" ("(1TYPEREP/(1'(_')))") wenzelm@35115: wenzelm@35115: parse_translation {* haftmann@26168: let haftmann@28335: fun typerep_tr (*"_TYPEREP"*) [ty] = wenzelm@35115: Syntax.const @{const_syntax typerep} $ wenzelm@35115: (Syntax.const @{syntax_const "_constrain"} $ Syntax.const @{const_syntax "TYPE"} $ wenzelm@35363: (Syntax.const @{type_syntax itself} $ ty)) haftmann@28335: | typerep_tr (*"_TYPEREP"*) ts = raise TERM ("typerep_tr", ts); wenzelm@35115: in [(@{syntax_const "_TYPEREP"}, typerep_tr)] end wenzelm@35115: *} wenzelm@35115: wenzelm@35115: typed_print_translation {* wenzelm@35115: let wenzelm@35363: fun typerep_tr' show_sorts (*"typerep"*) wenzelm@35430: (Type (@{type_name fun}, [Type (@{type_name itself}, [T]), _])) wenzelm@35363: (Const (@{const_syntax TYPE}, _) :: ts) = wenzelm@35115: Term.list_comb wenzelm@35115: (Syntax.const @{syntax_const "_TYPEREP"} $ Syntax.term_of_typ show_sorts T, ts) haftmann@28335: | typerep_tr' _ T ts = raise Match; wenzelm@35115: in [(@{const_syntax typerep}, typerep_tr')] end haftmann@26168: *} haftmann@26168: haftmann@31137: setup {* haftmann@31137: let haftmann@26168: haftmann@31137: fun add_typerep 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@31205: val rhs = @{term Typerep} $ HOLogic.mk_literal tyco haftmann@31137: $ HOLogic.mk_list @{typ typerep} (map (HOLogic.mk_typerep o TFree) vs); haftmann@26168: val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); haftmann@26168: in haftmann@26168: thy wenzelm@33553: |> Theory_Target.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@31137: |> Class.prove_instantiation_exit (K (Class.intro_classes_tac [])) haftmann@26168: end; haftmann@26168: haftmann@31137: fun ensure_typerep tyco thy = if not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep}) haftmann@31137: andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort type} haftmann@31137: then add_typerep tyco thy else thy; haftmann@31137: haftmann@31137: in haftmann@26168: haftmann@31137: add_typerep @{type_name fun} haftmann@31723: #> Typedef.interpretation ensure_typerep haftmann@35299: #> Code.datatype_interpretation (ensure_typerep o fst) haftmann@35299: #> Code.abstype_interpretation (ensure_typerep o fst) haftmann@26168: haftmann@31137: end 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@31031: (Eval "Term.typ") haftmann@26168: haftmann@28335: code_const Typerep haftmann@31031: (Eval "Term.Type/ (_, _)") haftmann@26168: haftmann@31031: code_reserved Eval Term haftmann@26168: wenzelm@36176: hide_const (open) typerep Typerep haftmann@26168: haftmann@26168: end