Add very deep interpretation

Use metalogic to generate meta term anti-quotations

The idea is for the Very_Deep_Interpretation
to source the shallow material,
and then update the checking and elaboration functions
of the term anti-quotations.
To achieve this, the mechanism of removing and readding the notations
(mixfixes) of the term-antiquotations after the metalogic
is sourced is used.

Example:

With shallow:

datatype "typ" = ISA_typ string  ("@{typ _}")

Generate a datatype whose Constructor ISA_typ has
the notation @{typ ...}.

You get:

find_consts name:"ISA_typ"

find_consts
  name: "ISA_typ"

found 1 constant(s):
  Shallow_Interpretation.typ.ISA_typ :: "char list ⇒ typ"

With Deep:

no_notation "ISA_typ" ("@{typ _}")

consts ISA_typ :: "string ⇒ typ" ("@{typ _}")

The notation is removed and then added to the new ISA_typ constant.

You get:

find_consts name:"ISA_typ"

find_consts
  name: "ISA_typ"

found 2 constant(s):
  Deep_Interpretation.ISA_typ :: "char list ⇒ Core.typ"
  Shallow_Interpretation.typ.ISA_typ :: "char list ⇒ Shallow_Interpretation.typ"

But only the Deep_Interpretation constant has the notation (mixfix).

Then new interpretation of term anti-quotations is available
for the user.
This commit is contained in:
Nicolas Méric 2023-01-16 15:52:23 +01:00
parent 8513f7d267
commit 756f2b66f1
8 changed files with 652 additions and 2 deletions

View File

@ -15,5 +15,3 @@ session "Isabelle_DOF" = "Functional-Automata" +
"DOF/Isa_DOF"
"ontologies/ontologies"
export_classpath

View File

@ -1 +1,2 @@
tests
very_deep_interpretation

View File

@ -0,0 +1,6 @@
session "Isabelle_Very_Deep_DOF" = Isabelle_DOF +
options [timeout = 0, record_proofs = 2]
sessions
"Metalogic_ProofChecker"
theories
"Very_Deep_DOF"

View File

@ -0,0 +1 @@
tests

View File

@ -0,0 +1,20 @@
theory Very_Deep_DOF
imports "Isabelle_Very_Deep_DOF.Very_Deep_Interpretation"
begin
(* tests *)
term "@{typ ''int => int''}"
term "@{term ''Bound 0''}"
term "@{thm ''refl''}"
term "@{docitem ''<doc_ref>''}"
ML\<open> @{term "@{docitem ''<doc_ref>''}"}\<close>
term "@{typ \<open>int \<Rightarrow> int\<close>}"
term "@{term \<open>\<forall>x. P x \<longrightarrow> Q\<close>}"
term "@{thm \<open>refl\<close>}"
term "@{docitem \<open>doc_ref\<close>}"
ML\<open> @{term "@{docitem \<open>doc_ref\<close>}"}\<close>
(**)
end

View File

@ -0,0 +1,315 @@
theory Very_Deep_Interpretation
imports "Isabelle_DOF.Isa_COL"
Metalogic_ProofChecker.ProofTerm
begin
subsection\<open> Syntax \<close>
\<comment> \<open>and others in the future : file, http, thy, ...\<close>
(* Delete shallow interpretation notations (mixfixes) of the term anti-quotations,
so we can use them for the deep interpretation *)
no_notation "ISA_typ" ("@{typ _}")
no_notation "ISA_term" ("@{term _}")
no_notation "ISA_thm" ("@{thm _}")
no_notation "ISA_file" ("@{file _}")
no_notation "ISA_thy" ("@{thy _}")
no_notation "ISA_docitem" ("@{docitem _}")
no_notation "ISA_docitem_attr" ("@{docitemattr (_) :: (_)}")
no_notation "ISA_trace_attribute" ("@{trace-attribute _}")
consts ISA_typ :: "string \<Rightarrow> typ" ("@{typ _}")
consts ISA_term :: "string \<Rightarrow> term" ("@{term _}")
consts ISA_term_repr :: "string \<Rightarrow> term" ("@{termrepr _}")
datatype "thm" = ISA_thm string ("@{thm _}") | Thm_content ("proof":proofterm)
datatype "thms_of" = ISA_thms_of string ("@{thms-of _}")
datatype "file" = ISA_file string ("@{file _}")
datatype "thy" = ISA_thy string ("@{thy _}")
consts ISA_docitem :: "string \<Rightarrow> 'a" ("@{docitem _}")
datatype "docitem_attr" = ISA_docitem_attr string string ("@{docitemattr (_) :: (_)}")
consts ISA_trace_attribute :: "string \<Rightarrow> (string * string) list" ("@{trace-attribute _}")
subsection\<open> Semantics \<close>
ML\<open>
structure ISA_core =
struct
fun check_path check_file ctxt dir (name, pos) =
let
val _ = Context_Position.report ctxt pos (Markup.language_path true); (* TODO: pos should be
"lifted" to
type source *)
val path = Path.append dir (Path.explode name) handle ERROR msg => ISA_core.err msg pos;
val _ = Path.expand path handle ERROR msg => ISA_core.err msg pos;
val _ = Context_Position.report ctxt pos (Markup.path (Path.implode_symbolic path));
val _ =
(case check_file of
NONE => path
| SOME check => (check path handle ERROR msg => ISA_core.err msg pos));
in path end;
fun ML_isa_antiq check_file thy (name, _, pos) =
let val path = check_path check_file (Proof_Context.init_global thy) Path.current (name, pos);
in "Path.explode " ^ ML_Syntax.print_string (Path.implode path) end;
fun ML_isa_check_generic check thy (term, pos) =
let val name = (HOLogic.dest_string term
handle TERM(_,[t]) => error ("wrong term format: must be string constant: "
^ Syntax.string_of_term_global thy t ))
val _ = check thy (name,pos)
in SOME term end;
fun check_identity _ (term, _, _) _ = SOME term
fun ML_isa_check_typ thy (term, _, pos) _ =
let fun check thy (name, _) = let val ctxt = (Proof_Context.init_global thy)
in (Syntax.check_typ ctxt o Syntax.parse_typ ctxt) name end
in ML_isa_check_generic check thy (term, pos) end
fun ML_isa_check_term thy (term, _, pos) _ =
let fun check thy (name, _) = let val ctxt = (Proof_Context.init_global thy)
in (Syntax.check_term ctxt o Syntax.parse_term ctxt) name end
in ML_isa_check_generic check thy (term, pos) end
fun ML_isa_check_thm thy (term, _, pos) _ =
(* this works for long-names only *)
let fun check thy (name, _) = case Proof_Context.lookup_fact (Proof_Context.init_global thy) name of
NONE => ISA_core.err ("No Theorem:" ^name) pos
| SOME X => X
in ML_isa_check_generic check thy (term, pos) end
fun ML_isa_check_file thy (term, _, pos) _ =
let fun check thy (name, pos) = check_path (SOME File.check_file)
(Proof_Context.init_global thy)
(Path.current)
(name, pos);
in ML_isa_check_generic check thy (term, pos) end;
fun ML_isa_id thy (term,pos) = SOME term
fun ML_isa_check_docitem thy (term, req_ty, pos) _ =
let fun check thy (name, _) s =
if DOF_core.is_declared_oid_global name thy
then case DOF_core.get_object_global name thy of
NONE => warning("oid declared, but not yet defined --- "^
" type-check incomplete")
| SOME {pos=pos_decl,cid,id,...} =>
let val ctxt = (Proof_Context.init_global thy)
val req_class = case req_ty of
\<^Type>\<open>fun _ T\<close> => DOF_core.typ_to_cid T
| _ => error("can not infer type for: "^ name)
in if cid <> DOF_core.default_cid
andalso not(DOF_core.is_subclass ctxt cid req_class)
then error("reference ontologically inconsistent: "
^cid^" vs. "^req_class^ Position.here pos_decl)
else ()
end
else ISA_core.err ("faulty reference to docitem: "^name) pos
in ML_isa_check_generic check thy (term, pos) end
fun ML_isa_check_trace_attribute thy (term, _, pos) s =
let
fun check thy (name, _) =
case DOF_core.get_object_global name thy of
NONE => ISA_core.err ("No class instance: " ^ name) pos
| SOME(_) => ()
in ML_isa_check_generic check thy (term, pos) end
fun ML_isa_elaborate_generic (_:theory) isa_name ty term_option _ =
case term_option of
NONE => error("Wrong term option. You must use a defined term")
| SOME term => Const (isa_name, ty) $ term
fun reify_typ (Type (s, typ_list)) =
\<^Const>\<open>Ty\<close> $ HOLogic.mk_literal s $ HOLogic.mk_list \<^Type>\<open>typ\<close> (map reify_typ typ_list)
| reify_typ (TFree (name, sort)) =
\<^Const>\<open>Tv\<close> $(\<^Const>\<open>Free\<close> $ HOLogic.mk_literal name)
$ (HOLogic.mk_set \<^typ>\<open>class\<close> (map HOLogic.mk_literal sort))
| reify_typ (TVar (indexname, sort)) =
let val (name, index_value) = indexname
in \<^Const>\<open>Tv\<close>
$ (\<^Const>\<open>Var\<close>
$ HOLogic.mk_prod (HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\<open>int\<close> index_value))
$ (HOLogic.mk_set \<^typ>\<open>class\<close> (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>\<open>Ct\<close> $ HOLogic.mk_literal name $ reify_typ typ
| reify_term (Free (name, typ)) =
\<^Const>\<open>Fv\<close> $ (\<^Const>\<open>Free\<close> $ HOLogic.mk_literal name) $ reify_typ typ
| reify_term (Var (indexname, typ)) =
let val (name, index_value) = indexname
in \<^Const>\<open>Fv\<close>
$ (\<^Const>\<open>Var\<close>
$ HOLogic.mk_prod (HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\<open>int\<close> index_value))
$ reify_typ typ end
| reify_term (Bound i) = \<^Const>\<open>Bv\<close> $ HOLogic.mk_nat i
| reify_term (Abs (_, typ, term)) = \<^Const>\<open>Abs\<close> $ reify_typ typ $ reify_term term
| reify_term (Term.$ (t1, t2)) = \<^Const>\<open>App\<close> $ 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>\<open>PBound\<close> $ (HOLogic.mk_nat i)
| reify_proofterm (Abst (_, typ_option, proof)) =
\<^Const>\<open>Abst\<close> $ reify_typ (the typ_option) $ reify_proofterm proof
| reify_proofterm (AbsP (_, term_option, proof)) =
\<^Const>\<open>AbsP\<close> $ reify_term (the term_option) $ reify_proofterm proof
| reify_proofterm (op % (proof, term_option)) =
\<^Const>\<open>Appt\<close> $ reify_proofterm proof $ reify_term (the term_option)
| reify_proofterm (op %% (proof1, proof2)) =
\<^Const>\<open>AppP\<close> $ reify_proofterm proof1 $ reify_proofterm proof2
| reify_proofterm (Hyp term) = \<^Const>\<open>Hyp\<close> $ (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>\<open>Var\<close>
$ HOLogic.mk_prod
(HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\<open>int\<close> index_value)
, HOLogic.mk_set \<^typ>\<open>class\<close> (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>\<open>PAxm\<close> $ reify_term term $ meta_typ_list end
| reify_proofterm (PClass (typ, class)) =
\<^Const>\<open>OfClass\<close> $ 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>\<open>Var\<close>
$ HOLogic.mk_prod
(HOLogic.mk_literal name, HOLogic.mk_number \<^Type>\<open>int\<close> index_value)
, HOLogic.mk_set \<^typ>\<open>class\<close> (map HOLogic.mk_literal sort))) tvars
val meta_typ_list =
HOLogic.mk_list \<^typ>\<open>tyinst\<close> (map2 (fn x => fn y => HOLogic.mk_prod (x, y))
meta_tvars (map reify_typ (the types)))
in \<^Const>\<open>PAxm\<close> $ 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>\<open>Thm_content\<close> $ reify_proofterm prf end*)
(*in \<^Const>\<open>Thm_content\<close> $ reify_proofterm proof end*)
in \<^Const>\<open>Thm_content\<close> $ 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>\<open>thm\<close> (map (fn proof => \<^Const>\<open>Thm_content\<close> $ reify_proofterm proof) all_proofs) end*)
in HOLogic.mk_list \<^typ>\<open>string\<close> (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>\<open>Pair \<^typ>\<open>doc_class rexp\<close> \<^typ>\<open>string\<close>\<close>
$ (\<^Const>\<open>Atom \<^typ>\<open>doc_class\<close>\<close> $ (\<^Const>\<open>mk\<close> $ s)) $ S) =
let val s' = DOF_core.read_cid (Proof_Context.init_global thy) (HOLogic.dest_string s)
in \<^Const>\<open>Pair \<^typ>\<open>string\<close> \<^typ>\<open>string\<close>\<close> $ HOLogic.mk_string s' $ S end
val traces' = map conv (HOLogic.dest_list traces)
in HOLogic.mk_list \<^Type>\<open>prod \<^typ>\<open>string\<close> \<^typ>\<open>string\<close>\<close> traces' end
(* utilities *)
fun property_list_dest ctxt X =
map (fn \<^Const_>\<open>ISA_term for s\<close> => HOLogic.dest_string s
|\<^Const_>\<open>ISA_term_repr for s\<close> => holstring_to_bstring ctxt (HOLogic.dest_string s))
(HOLogic.dest_list X)
end; (* struct *)
\<close>
ML\<open>
val ty1 = ISA_core.reify_typ @{typ "int"}
val ty2 = ISA_core.reify_typ @{typ "int \<Rightarrow> bool"}
val ty3 = ISA_core.reify_typ @{typ "prop"}
val ty4 = ISA_core.reify_typ @{typ "'a list"}
\<close>
ML\<open>
val t1 = ISA_core.reify_term @{term "1::int"}
val t2 = ISA_core.reify_term @{term "\<lambda>x. x = 1"}
val t3 = ISA_core.reify_term @{term "[2, 3::int]"}
\<close>
subsection\<open> Isar - Setup\<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.typ",
{check=ISA_core.ML_isa_check_typ, elaborate=ISA_core.ML_isa_elaborate_typ}) \<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.term",
{check=ISA_core.ML_isa_check_term, elaborate=ISA_core.ML_isa_elaborate_term}) \<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.term_repr",
{check=ISA_core.check_identity, elaborate=ISA_core.ML_isa_elaborate_generic}) \<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.thm.thm",
{check=ISA_core.ML_isa_check_thm, elaborate=ISA_core.ML_isa_elaborate_thm}) \<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.thms_of.thms_of",
{check=ISA_core.ML_isa_check_thm, elaborate=ISA_core.ML_isa_elaborate_thms_of}) \<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.file.file",
{check=ISA_core.ML_isa_check_file, elaborate=ISA_core.ML_isa_elaborate_generic}) \<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.docitem",
{check=ISA_core.ML_isa_check_docitem, elaborate=ISA_core.ML_isa_elaborate_generic}) \<close>
setup\<open>DOF_core.update_isa_global("Very_Deep_Interpretation.trace_attribute",
{check=ISA_core.ML_isa_check_trace_attribute, elaborate=ISA_core.ML_isa_elaborate_trace_attribute}) \<close>
end

View File

@ -0,0 +1,3 @@
session "Isabelle_Very_Deep_DOF-tests" = "Isabelle_Very_Deep_DOF" +
theories
"Reification_Test"

View File

@ -0,0 +1,306 @@
theory Reification_Test
imports "../DOF_Deep"
begin
ML\<open>
val ty1 = ISA_core.reify_typ @{typ "int"}
val ty2 = ISA_core.reify_typ @{typ "int \<Rightarrow> bool"}
val ty3 = ISA_core.reify_typ @{typ "prop"}
val ty4 = ISA_core.reify_typ @{typ "'a list"}
\<close>
term*\<open>@{typ \<open>int\<close>}\<close>
value*\<open>@{typ \<open>int\<close>}\<close>
value*\<open>@{typ \<open>int \<Rightarrow> bool\<close>}\<close>
term*\<open>@{typ \<open>prop\<close>}\<close>
value*\<open>@{typ \<open>prop\<close>}\<close>
term*\<open>@{typ \<open>'a list\<close>}\<close>
value*\<open>@{typ \<open>'a list\<close>}\<close>
ML\<open>
val t1 = ISA_core.reify_term @{term "1::int"}
val t2 = ISA_core.reify_term @{term "\<lambda>x. x = 1"}
val t3 = ISA_core.reify_term @{term "[2, 3::int]"}
\<close>
term*\<open>@{term \<open>1::int\<close>}\<close>
value*\<open>@{term \<open>1::int\<close>}\<close>
term*\<open>@{term \<open>\<lambda>x. x = 1\<close>}\<close>
value*\<open>@{term \<open>\<lambda>x. x = 1\<close>}\<close>
term*\<open>@{term \<open>[2, 3::int]\<close>}\<close>
value*\<open>@{term \<open>[2, 3::int]\<close>}\<close>
prf refl
full_prf refl
term*\<open>@{thm \<open>HOL.refl\<close>}\<close>
value*\<open>proof @{thm \<open>HOL.refl\<close>}\<close>
value*\<open>depth (proof @{thm \<open>HOL.refl\<close>})\<close>
value*\<open>size (proof @{thm \<open>HOL.refl\<close>})\<close>
value*\<open>fv_Proof (proof @{thm \<open>HOL.refl\<close>})\<close>
term*\<open>@{thms-of \<open>HOL.refl\<close>}\<close>
value*\<open>@{thms-of \<open>HOL.refl\<close>}\<close>
ML\<open>
val t_schematic = TVar(("'a",0), [])
val t = @{term "Tv (Var (STR '''a'', 0)) {}"}
val rt_schematic = ISA_core.reify_typ t_schematic
val true = rt_schematic = t
\<close>
lemma test : "AAA \<and> BBB \<longrightarrow> BBB \<and> AAA"
by auto
lemma test_subst : "\<lbrakk>x = f x; odd(f x)\<rbrakk> \<Longrightarrow> odd x"
by (erule ssubst)
inductive_set even :: "int set" where
zero[intro!]: "0 \<in> even" |
plus[intro!]: "n \<in> even \<Longrightarrow> n+2 \<in> even " |
min[intro!]: "n \<in> even \<Longrightarrow> n-2 \<in> even "
lemma a : "2+2=4" by simp
lemma b : "(0::int)+2=2" by simp
lemma test_subst_2 : "4 \<in> even"
apply (subst a[symmetric])
apply (rule plus)
apply (subst b[symmetric])
apply (rule plus)
by auto
(*lemma "\<lbrakk>P x y z; Suc x < y\<rbrakk> \<Longrightarrow> f z = x * y"
(*using [[simp_trace]]*)
(*apply (simp add: mult.commute)*)
apply (subst mult.commute)
apply (rule mult.commute [THEN ssubst])*)
datatype 'a seq = Empty | Seq 'a "'a seq"
find_consts name:"Reification_Test*seq*"
fun conc :: "'a seq \<Rightarrow> 'a seq \<Rightarrow> 'a seq"
where
c1 : "conc Empty ys = ys"
| c2 : "conc (Seq x xs) ys = Seq x (conc xs ys)"
lemma identity_conc : "conc xs Empty = xs"
using [[simp_trace]]
using[[simp_trace_depth_limit=8]]
using [[unify_trace_simp]]
using[[unify_trace_types]]
using [[unify_trace_bound=0]]
(* using [[simp_trace_new depth=10]] *)
apply (induct xs)
apply (subst c1)
apply (rule refl)
apply (subst c2)
apply (rule_tac s="xs" and P="\<lambda>X. Seq x1 X = Seq x1 xs" in subst)
apply (rule sym)
apply assumption
apply (rule refl)
done
lemma conj_test : "A \<and> B \<and> C \<longrightarrow> B \<and> A"
apply (rule impI)
apply (rule conjI)
apply (drule conjunct2)
apply (drule conjunct1)
apply assumption
apply (drule conjunct1)
apply assumption
done
lemma test2 : "AAA \<and> BBB \<Longrightarrow> BBB \<and> AAA"
by auto
lemma test3: "AAAAA \<and> BBBBB \<longrightarrow> BBBBB \<and> AAAAA"
proof
assume "AAAAA \<and> BBBBB"
then obtain BBBBB and AAAAA ..
then show "BBBBB \<and> AAAAA" ..
qed
lemma test4:
assumes "(AAA \<and> BBB)"
shows "BBB \<and> AAA"
apply (insert assms)
by auto
declare[[show_sorts]]
declare[[ML_print_depth = 20]]
ML\<open>
val full = true
val thm = @{thm "test"}
val hyps = Thm.hyps_of thm
val prems = Thm.prems_of thm
val reconstruct_proof = Thm.reconstruct_proof_of thm
val standard_proof = Proof_Syntax.standard_proof_of
{full = full, expand_name = Thm.expand_name thm} thm
val term_of_proof = Proof_Syntax.term_of_proof standard_proof
\<close>
ML\<open> (*See: *) \<^file>\<open>~~/src/HOL/Proofs/ex/Proof_Terms.thy\<close>\<close>
ML\<open>
val thm = @{thm test};
(*proof body with digest*)
val body = Proofterm.strip_thm_body (Thm.proof_body_of thm);
(*proof term only*)
val prf = Proofterm.proof_of body;
(*clean output*)
Pretty.writeln (Proof_Syntax.pretty_proof \<^context> prf);
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> false thm);
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm);
(*all theorems used in the graph of nested proofs*)
val all_thms =
Proofterm.fold_body_thms
(fn {name, ...} => insert (op =) name) [body] [];
\<close>
prf test
full_prf test
term*\<open>@{thm \<open>Reification_Test.test\<close>}\<close>
value*\<open>@{thm \<open>Reification_Test.test\<close>}\<close>
term*\<open>@{thms-of \<open>Reification_Test.test\<close>}\<close>
value*\<open>@{thms-of \<open>Reification_Test.test\<close>}\<close>
ML\<open> (*See: *) \<^file>\<open>~~/src/HOL/Proofs/ex/Proof_Terms.thy\<close>\<close>
ML\<open>
val thm = @{thm test4};
(*proof body with digest*)
val body = Proofterm.strip_thm_body (Thm.proof_body_of thm);
(*proof term only*)
val prf = Proofterm.proof_of body;
(*clean output*)
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> false thm);
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm);
(*all theorems used in the graph of nested proofs*)
val all_thms =
Proofterm.fold_body_thms
(fn {name, ...} => insert (op =) name) [body] [];
\<close>
ML\<open> (*See: *) \<^file>\<open>~~/src/HOL/Proofs/ex/Proof_Terms.thy\<close>\<close>
ML\<open>
val thm = @{thm test2};
(*proof body with digest*)
val body = Proofterm.strip_thm_body (Thm.proof_body_of thm);
(*proof term only*)
val prf = Proofterm.proof_of body;
(*clean output*)
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> false thm);
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm);
(*all theorems used in the graph of nested proofs*)
val all_thms =
Proofterm.fold_body_thms
(fn {name, ...} => insert (op =) name) [body] [];
\<close>
prf test2
full_prf test2
term*\<open>@{thm \<open>Reification_Test.test2\<close>}\<close>
value*\<open>proof @{thm \<open>Reification_Test.test2\<close>}\<close>
ML\<open> (*See: *) \<^file>\<open>~~/src/HOL/Proofs/ex/Proof_Terms.thy\<close>\<close>
ML\<open>
val thm = @{thm test3};
(*proof body with digest*)
val body = Proofterm.strip_thm_body (Thm.proof_body_of thm);
(*proof term only*)
val prf = Proofterm.proof_of body;
(*clean output*)
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> false thm);
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm);
(*all theorems used in the graph of nested proofs*)
val all_thms =
Proofterm.fold_body_thms
(fn {name, ...} => insert (op =) name) [body] [];
\<close>
prf test2
full_prf test2
term*\<open>@{thm \<open>Reification_Test.test3\<close>}\<close>
value*\<open>@{thm \<open>Reification_Test.test3\<close>}\<close>
ML\<open> (*See: *) \<^file>\<open>~~/src/HOL/Proofs/ex/Proof_Terms.thy\<close>\<close>
ML\<open>
val thm = @{thm Pure.symmetric};
(*proof body with digest*)
val body = Proofterm.strip_thm_body (Thm.proof_body_of thm);
(*proof term only*)
val prf = Proofterm.proof_of body;
(*clean output*)
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> false thm);
Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm);
(*all theorems used in the graph of nested proofs*)
val all_thms =
Proofterm.fold_body_thms
(fn {name, ...} => insert (op =) name) [body] [];
\<close>
prf symmetric
full_prf symmetric
term*\<open>@{thm \<open>Pure.symmetric\<close>}\<close>
value*\<open>@{thm \<open>Pure.symmetric\<close>}\<close>
ML\<open>
val full = true
val thm = @{thm "Groups.minus_class.super"}
val standard_proof = Proof_Syntax.standard_proof_of
{full = full, expand_name = Thm.expand_name thm} thm
val term_of_proof = Proof_Syntax.term_of_proof standard_proof
\<close>
ML\<open>
val thm = Proof_Context.get_thm \<^context> "Groups.minus_class.super"
val prop = Thm.prop_of thm
val proof = Thm.proof_of thm
\<close>
prf Groups.minus_class.super
full_prf Groups.minus_class.super
term*\<open>@{thm \<open>Groups.minus_class.super\<close>}\<close>
value*\<open>@{thm \<open>Groups.minus_class.super\<close>}\<close>
(*ML\<open>
val full = true
val thm = @{thm "Homotopy.starlike_imp_contractible"}
val standard_proof = Proof_Syntax.standard_proof_of
{full = full, expand_name = Thm.expand_name thm} thm
val term_of_proof = Proof_Syntax.term_of_proof standard_proof
\<close>
ML\<open>
val thm = Proof_Context.get_thm \<^context> "Homotopy.starlike_imp_contractible"
val prop = Thm.prop_of thm
val proof = Thm.proof_of thm
\<close>
prf Homotopy.starlike_imp_contractible
full_prf Homotopy.starlike_imp_contractible
term*\<open>@{thm \<open>Homotopy.starlike_imp_contractible\<close>}\<close>
value*\<open>@{thm \<open>Homotopy.starlike_imp_contractible\<close>}\<close>*)
end