theory Very_Deep_Interpretation imports "Isabelle_DOF.Isa_COL" Metalogic_ProofChecker.ProofTerm begin subsection\ Syntax \ \ \and others in the future : file, http, thy, ...\ (* Delete shallow interpretation notations (mixfixes) of the term anti-quotations, so we can use them for the deep interpretation *) no_notation "Isabelle_DOF_typ" ("@{typ _}") no_notation "Isabelle_DOF_term" ("@{term _}") no_notation "Isabelle_DOF_thm" ("@{thm _}") no_notation "Isabelle_DOF_file" ("@{file _}") no_notation "Isabelle_DOF_thy" ("@{thy _}") no_notation "Isabelle_DOF_docitem" ("@{docitem _}") no_notation "Isabelle_DOF_docitem_attr" ("@{docitemattr (_) :: (_)}") no_notation "Isabelle_DOF_trace_attribute" ("@{trace'_-attribute _}") consts Isabelle_DOF_typ :: "string \ typ" ("@{typ _}") consts Isabelle_DOF_term :: "string \ term" ("@{term _}") datatype "thm" = Isabelle_DOF_thm string ("@{thm _}") | Thm_content ("proof":proofterm) datatype "thms_of" = Isabelle_DOF_thms_of string ("@{thms-of _}") datatype "file" = Isabelle_DOF_file string ("@{file _}") datatype "thy" = Isabelle_DOF_thy string ("@{thy _}") consts Isabelle_DOF_docitem :: "string \ 'a" ("@{docitem _}") datatype "docitem_attr" = Isabelle_DOF_docitem_attr string string ("@{docitemattr (_) :: (_)}") consts Isabelle_DOF_trace_attribute :: "string \ (string * string) list" ("@{trace'_-attribute _}") subsection\ Semantics \ ML\ structure Meta_ISA_core = struct fun ML_isa_check_trace_attribute thy (term, _, pos) s = let val oid = (HOLogic.dest_string term handle TERM(_,[t]) => error ("wrong term format: must be string constant: " ^ Syntax.string_of_term_global thy t )) val _ = DOF_core.get_instance_global oid thy in SOME term end fun reify_typ (Type (s, typ_list)) = \<^Const>\Ty\ $ HOLogic.mk_literal s $ HOLogic.mk_list \<^Type>\typ\ (map reify_typ typ_list) | reify_typ (TFree (name, sort)) = \<^Const>\Tv\ $(\<^Const>\Free\ $ HOLogic.mk_literal name) $ (HOLogic.mk_set \<^typ>\class\ (map HOLogic.mk_literal sort)) | reify_typ (TVar (indexname, sort)) = let val (name, index_value) = indexname in \<^Const>\Tv\ $ (\<^Const>\Var\ $ HOLogic.mk_prod (HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\int\ index_value)) $ (HOLogic.mk_set \<^typ>\class\ (map HOLogic.mk_literal sort)) end fun ML_isa_elaborate_typ (thy:theory) _ _ term_option _ = case term_option of NONE => error("Wrong term option. You must use a defined term") | SOME term => let val typ_name = HOLogic.dest_string term val typ = Syntax.read_typ_global thy typ_name in reify_typ typ end fun reify_term (Const (name, typ)) =\<^Const>\Ct\ $ HOLogic.mk_literal name $ reify_typ typ | reify_term (Free (name, typ)) = \<^Const>\Fv\ $ (\<^Const>\Free\ $ HOLogic.mk_literal name) $ reify_typ typ | reify_term (Var (indexname, typ)) = let val (name, index_value) = indexname in \<^Const>\Fv\ $ (\<^Const>\Var\ $ HOLogic.mk_prod (HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\int\ index_value)) $ reify_typ typ end | reify_term (Bound i) = \<^Const>\Bv\ $ HOLogic.mk_nat i | reify_term (Abs (_, typ, term)) = \<^Const>\Abs\ $ reify_typ typ $ reify_term term | reify_term (Term.$ (t1, t2)) = \<^Const>\App\ $ reify_term t1 $ reify_term t2 fun ML_isa_elaborate_term (thy:theory) _ _ term_option _ = case term_option of NONE => error("Wrong term option. You must use a defined term") | SOME term => let val term_name = HOLogic.dest_string term val term = Syntax.read_term_global thy term_name in reify_term term end fun reify_proofterm (PBound i) =\<^Const>\PBound\ $ (HOLogic.mk_nat i) | reify_proofterm (Abst (_, typ_option, proof)) = \<^Const>\Abst\ $ reify_typ (the typ_option) $ reify_proofterm proof | reify_proofterm (AbsP (_, term_option, proof)) = \<^Const>\AbsP\ $ reify_term (the term_option) $ reify_proofterm proof | reify_proofterm (op % (proof, term_option)) = \<^Const>\Appt\ $ reify_proofterm proof $ reify_term (the term_option) | reify_proofterm (op %% (proof1, proof2)) = \<^Const>\AppP\ $ reify_proofterm proof1 $ reify_proofterm proof2 | reify_proofterm (Hyp term) = \<^Const>\Hyp\ $ (reify_term term) | reify_proofterm (PAxm (_, term, typ_list_option)) = let val tvars = rev (Term.add_tvars term []) val meta_tvars = map (fn ((name, index_value), sort) => HOLogic.mk_prod (\<^Const>\Var\ $ HOLogic.mk_prod (HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\int\ index_value) , HOLogic.mk_set \<^typ>\class\ (map HOLogic.mk_literal sort))) tvars val meta_typ_list = HOLogic.mk_list @{typ "tyinst"} (map2 (fn x => fn y => HOLogic.mk_prod (x, y)) meta_tvars (map reify_typ (the typ_list_option))) in \<^Const>\PAxm\ $ reify_term term $ meta_typ_list end | reify_proofterm (PClass (typ, class)) = \<^Const>\OfClass\ $ reify_typ typ $ HOLogic.mk_literal class | reify_proofterm (PThm ({prop = prop, types = types, ...}, _)) = let val tvars = rev (Term.add_tvars prop []) val meta_tvars = map (fn ((name, index_value), sort) => HOLogic.mk_prod (\<^Const>\Var\ $ HOLogic.mk_prod (HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\int\ index_value) , HOLogic.mk_set \<^typ>\class\ (map HOLogic.mk_literal sort))) tvars val meta_typ_list = HOLogic.mk_list \<^typ>\tyinst\ (map2 (fn x => fn y => HOLogic.mk_prod (x, y)) meta_tvars (map reify_typ (the types))) in \<^Const>\PAxm\ $ reify_term prop $ meta_typ_list end fun ML_isa_elaborate_thm (thy:theory) _ _ term_option pos = case term_option of NONE => ISA_core.err ("Malformed term annotation") pos | SOME term => let val thm_name = HOLogic.dest_string term val _ = writeln ("In ML_isa_elaborate_thm thm_name: " ^ \<^make_string> thm_name) val thm = Proof_Context.get_thm (Proof_Context.init_global thy) thm_name val _ = writeln ("In ML_isa_elaborate_thm thm: " ^ \<^make_string> thm) val body = Proofterm.strip_thm_body (Thm.proof_body_of thm); val prf = Proofterm.proof_of body; (* Proof_Syntax.standard_proof_of reconstructs the proof and seems to rewrite the option arguments (with a value NONE) of the proof datatype constructors, at least for PAxm, with "SOME (typ/term)", allowing us the use the projection function "the". Maybe the function can deal with all the option types of the proof datatype constructors *) val proof = Proof_Syntax.standard_proof_of {full = true, expand_name = Thm.expand_name thm} thm val _ = writeln ("In ML_isa_elaborate_thm proof: " ^ \<^make_string> proof) (* After a small discussion with Simon Roßkopf, It seems preferable to use Thm.reconstruct_proof_of instead of Proof_Syntax.standard_proof_of whose operation is not well known. Thm.reconstruct_proof_of seems sufficient to have a reifiable PAxm in the metalogic. *) val proof' = Thm.reconstruct_proof_of thm (*in \<^Const>\Thm_content\ $ reify_proofterm prf end*) (*in \<^Const>\Thm_content\ $ reify_proofterm proof end*) in \<^Const>\Thm_content\ $ reify_proofterm proof end fun ML_isa_elaborate_thms_of (thy:theory) _ _ term_option pos = case term_option of NONE => ISA_core.err ("Malformed term annotation") pos | SOME term => let val thm_name = HOLogic.dest_string term val thm = Proof_Context.get_thm (Proof_Context.init_global thy) thm_name val body = Proofterm.strip_thm_body (Thm.proof_body_of thm) val all_thms_name = Proofterm.fold_body_thms (fn {name, ...} => insert (op =) name) [body] [] (*val all_thms = map (Proof_Context.get_thm (Proof_Context.init_global thy)) all_thms_name*) (*val all_proofs = map (Proof_Syntax.standard_proof_of {full = true, expand_name = Thm.expand_name thm}) all_thms*) (*in HOLogic.mk_list \<^Type>\thm\ (map (fn proof => \<^Const>\Thm_content\ $ reify_proofterm proof) all_proofs) end*) in HOLogic.mk_list \<^typ>\string\ (map HOLogic.mk_string all_thms_name) end fun ML_isa_elaborate_trace_attribute (thy:theory) _ _ term_option pos = case term_option of NONE => ISA_core.err ("Malformed term annotation") pos | SOME term => let val oid = HOLogic.dest_string term val traces = ISA_core.compute_attr_access (Context.Theory thy) "trace" oid NONE pos fun conv (\<^Const>\Pair \<^typ>\doc_class rexp\ \<^typ>\string\\ $ (\<^Const>\Atom \<^typ>\doc_class\\ $ (\<^Const>\mk\ $ s)) $ S) = let val s' = DOF_core.get_onto_class_name_global (HOLogic.dest_string s) thy in \<^Const>\Pair \<^typ>\string\ \<^typ>\string\\ $ HOLogic.mk_string s' $ S end val traces' = map conv (HOLogic.dest_list traces) in HOLogic.mk_list \<^Type>\prod \<^typ>\string\ \<^typ>\string\\ traces' end end; (* struct *) \ ML\ val ty1 = Meta_ISA_core.reify_typ @{typ "int"} val ty2 = Meta_ISA_core.reify_typ @{typ "int \ bool"} val ty3 = Meta_ISA_core.reify_typ @{typ "prop"} val ty4 = Meta_ISA_core.reify_typ @{typ "'a list"} \ ML\ val t1 = Meta_ISA_core.reify_term @{term "1::int"} val t2 = Meta_ISA_core.reify_term @{term "\x. x = 1"} val t3 = Meta_ISA_core.reify_term @{term "[2, 3::int]"} \ subsection\ Isar - Setup\ (* Isa_transformers declaration for Isabelle_DOF term anti-quotations (typ, term, thm, etc.). They must be declared in the same theory file as the one of the declaration of Isabelle_DOF term anti-quotations !!! *) setup\ [(\<^type_name>\thm\, ISA_core.ML_isa_check_thm, Meta_ISA_core.ML_isa_elaborate_thm) , (\<^type_name>\thms_of\, ISA_core.ML_isa_check_thm, Meta_ISA_core.ML_isa_elaborate_thms_of) , (\<^type_name>\file\, ISA_core.ML_isa_check_file, ISA_core.ML_isa_elaborate_generic)] |> fold (fn (n, check, elaborate) => fn thy => let val ns = Sign.tsig_of thy |> Type.type_space val name = n val {pos, ...} = Name_Space.the_entry ns name val bname = Long_Name.base_name name val binding = Binding.make (bname, pos) |> Binding.prefix_name DOF_core.ISA_prefix |> Binding.prefix false bname in DOF_core.add_isa_transformer binding ((check, elaborate) |> DOF_core.make_isa_transformer) thy end) #> ([(\<^const_name>\Isabelle_DOF_typ\, ISA_core.ML_isa_check_typ, Meta_ISA_core.ML_isa_elaborate_typ) ,(\<^const_name>\Isabelle_DOF_term\, ISA_core.ML_isa_check_term, Meta_ISA_core.ML_isa_elaborate_term) ,(\<^const_name>\Isabelle_DOF_docitem\, ISA_core.ML_isa_check_docitem, ISA_core.ML_isa_elaborate_generic) ,(\<^const_name>\Isabelle_DOF_trace_attribute\, ISA_core.ML_isa_check_trace_attribute, ISA_core.ML_isa_elaborate_trace_attribute)] |> fold (fn (n, check, elaborate) => fn thy => let val ns = Sign.consts_of thy |> Consts.space_of val name = n val {pos, ...} = Name_Space.the_entry ns name val bname = Long_Name.base_name name val binding = Binding.make (bname, pos) in DOF_core.add_isa_transformer binding ((check, elaborate) |> DOF_core.make_isa_transformer) thy end)) \ end