theory Reification_Test imports "Isabelle_DOF-Proofs.Very_Deep_DOF" begin 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"} \ term*\@{typ \int\}\ value*\@{typ \int\}\ value*\@{typ \int \ bool\}\ term*\@{typ \prop\}\ value*\@{typ \prop\}\ term*\@{typ \'a list\}\ value*\@{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]"} \ term*\@{term \1::int\}\ value*\@{term \1::int\}\ term*\@{term \\x. x = 1\}\ value*\@{term \\x. x = 1\}\ term*\@{term \[2, 3::int]\}\ value*\@{term \[2, 3::int]\}\ prf refl full_prf refl term*\@{thm \HOL.refl\}\ value*\proof @{thm \HOL.refl\}\ value*\proof @{thm \HOL.refl\}\ value*\depth (proof @{thm \HOL.refl\})\ value*\size (proof @{thm \HOL.refl\})\ value*\fv_Proof (proof @{thm \HOL.refl\})\ term*\@{thms-of \HOL.refl\}\ value*\@{thms-of \HOL.refl\}\ ML\ val t_schematic = TVar(("'a",0), []) val t = @{term "Tv (Var (STR '''a'', 0)) {}"} val rt_schematic = Meta_ISA_core.reify_typ t_schematic val true = rt_schematic = t \ lemma test : "A \ B \ B \ A" by auto lemma test2 : "A \ B \ B \ A" by auto lemma test3: "A \ B \ B \ A" proof assume "A \ B" then obtain B and A .. then show "B \ A" .. qed lemma test4: assumes "(A \ B)" shows "B \ A" apply (insert assms) by auto lemma test_subst : "\x = f x; odd(f x)\ \ odd x" by (erule ssubst) inductive_set even' :: "nat set" where "0 \ even'" | "n \ even' \ (Suc (Suc n)) \ even'" find_theorems name:"even'.induct" (*lemma even_dvd : "n \ even' \ 2 dvd n" proof(induct n) case 0 then show ?case by simp next case (Suc n) then show ?case apply (simp add: dvd_def) apply (rule_tac x ="Suc k" in exI) apply clarify*) theorem "((A \ B) \ A) \ A" proof assume "(A \ B) \ A" show A proof (rule classical) assume "\ A" have "A \ B" proof assume A with \\ A\ show B by contradiction qed with \(A \ B) \ A\ show A .. qed qed (*lemma even_dvd : "n \ even' \ 2 dvd n" using [[simp_trace]] apply (induct n) apply (subst even_zero) apply(rule TrueI) apply(simp)*) lemma even_dvd : "n \ even' \ 2 dvd n" apply (erule even'.induct) apply (simp_all add: dvd_def) using [[simp_trace]] apply clarify find_theorems name:"_ = 2 * _" apply (rule_tac x ="Suc k" in exI) using [[simp_trace]] apply simp done (* lemma even_dvd : "n \ even' \ 2 dvd n" apply (induct_tac rule:even'.induct)*) inductive ev :: " nat \ bool " where ev0: " ev 0 " | evSS: " ev n \ ev (n + 2) " fun evn :: " nat \ bool " where " evn 0 = True " | " evn (Suc 0) = False " | " evn (Suc (Suc n)) = evn n " (*lemma assumes a: " ev (Suc(Suc m)) " shows" ev m " proof(induction "Suc (Suc m)" arbitrary: " m " rule: ev.induct)*) (*lemma " ev (Suc (Suc m)) \ ev m " proof(induction " Suc (Suc m) " arbitrary: " m " rule: ev.induct) case ev0 then show ?case sorry next case (evSS n) then show ?case sorry qed*) (* And neither of these can apply the induction *) (* lemma assumes a1: " ev n " and a2: " n = (Suc (Suc m)) " shows " ev m " proof (induction " n " arbitrary: " m " rule: ev.induct) lemma assumes a1: " n = (Suc (Suc m)) " and a2: "ev n " shows " ev m " proof (induction " n " arbitrary: " m " rule: ev.induct) *) (* But this one can ?! *) (* lemma assumes a1: " ev n " and a2: " n = (Suc (Suc m)) " shows " ev m " proof - from a1 and a2 show " ev m " proof (induction " n " arbitrary: " m " rule: ev.induct) case ev0 then show ?case by simp next case (evSS n) thus ?case by simp qed qed *) inductive_set even :: "int set" where zero[intro!]: "0 \ even" | plus[intro!]: "n \ even \ n+2 \ even " | min[intro!]: "n \ even \ n-2 \ even " lemma a : "2+2=4" by simp lemma b : "(0::int)+2=2" by simp lemma test_subst_2 : "4 \ even" apply (subst a[symmetric]) apply (rule plus) apply (subst b[symmetric]) apply (rule plus) apply (rule zero) done (*lemma "\P x y z; Suc x < y\ \ 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 \ 'a seq \ 'a seq" where c1 : "conc Empty ys = ys" | c2 : "conc (Seq x xs) ys = Seq x (conc xs ys)" lemma seq_not_eq : "Seq x xs \ xs" using [[simp_trace]] proof (induct xs arbitrary: x) case Empty show "Seq x Empty \ Empty" by simp next case (Seq y ys) show "Seq x (Seq y ys) \ Seq y ys" using \Seq y ys \ ys\ by simp qed 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="\X. Seq x1 X = Seq x1 xs" in subst) apply (rule sym) apply assumption apply (rule refl) done lemma imp_ex : "(\x. \y. P x y) \ (\y. \x. P x y)" using [[simp_trace]] using[[simp_trace_depth_limit=8]] apply (auto) done lemma length_0_conv' [iff]: "(length [] = 0)" apply (subst List.list.size(3)) apply (rule refl) done lemma cons_list : "a#xs = [a]@xs" using [[simp_trace]] apply (subst List.append.append_Cons) apply (subst List.append.append_Nil) apply (rule refl) done lemma replacement: "\ a = b; c = d \ \ f a c = f b d" apply (erule ssubst)+ apply (rule refl ) done lemma assoc_append : "k @ (l @ m) = (k @ l ) @ m" apply (induct_tac k ) apply (subst append_Nil )+ apply (rule refl ) apply (subst append_Cons) apply (subst append_Cons) apply (subst append_Cons) apply (rule_tac f ="Cons" in replacement) apply (rule refl) apply assumption done lemma length_cons : "length (xs @ ys) = length xs + length ys" using [[simp_trace]] apply (subst List.length_append) apply (rule refl) done lemma length_plus : "(length [a] + length xs = 0) = ([a] @ xs = [])" using [[simp_trace]] apply (subst List.list.size(4)) apply (subst List.list.size(3)) apply (subst Nat.add_Suc_right) apply (subst Groups.monoid_add_class.add.right_neutral) apply (subst Nat.plus_nat.add_Suc) apply (subst Groups.monoid_add_class.add.left_neutral) apply (subst Nat.old.nat.distinct(2)) by simp lemma empty_list : "(length [] = 0) = ([] = []) = True" using [[simp_trace]] by simp lemma TrueI: True using [[simp_trace]] unfolding True_def by (rule refl) lemma length_0_conv [iff]: "(length xs = 0) = (xs = [])" using [[simp_trace]] apply (induct xs) apply (subst List.list.size(3)) apply(subst HOL.simp_thms(6)) apply(subst HOL.simp_thms(6)) apply(rule refl) apply (subst cons_list) apply (subst(2) cons_list) apply (subst length_cons) apply (subst length_plus) apply (subst HOL.simp_thms(6)) apply (rule TrueI) done (*by (induct xs) auto*) find_theorems (50) name:"HOL.simp_thms" find_theorems (50) name:"List.list*size" find_theorems (50) name:"List.list*length" find_theorems "_ @ _" find_theorems (500) "List.length [] = 0" find_theorems (550) "length _ = length _ + length _" lemma identity_list : "xs @ [] = xs" using [[simp_trace]] using[[simp_trace_depth_limit=8]] using [[unify_trace_simp]] using[[unify_trace_types]] using [[unify_trace_bound=0]] apply (induct xs) apply (subst List.append_Nil2) apply (subst HOL.simp_thms(6)) apply(rule TrueI) apply (subst List.append_Nil2) apply (subst HOL.simp_thms(6)) apply(rule TrueI) done lemma identity_list' : "xs @ [] = 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 "length xs") apply (subst (asm) zero_reorient) apply(subst(asm) length_0_conv) apply (subst List.append_Nil2) apply (subst HOL.simp_thms(6)) apply (rule TrueI) apply (subst List.append_Nil2) apply (subst HOL.simp_thms(6)) apply (rule TrueI) done lemma conj_test : "A \ B \ C \ B \ A" apply (rule impI) apply (rule conjI) apply (drule conjunct2) apply (drule conjunct1) apply assumption apply (drule conjunct1) apply assumption done declare[[show_sorts]] declare[[ML_print_depth = 20]] ML\ 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 \ 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="\X. Seq x1 X = Seq x1 xs" in subst) apply (rule sym) apply assumption apply (rule refl) done declare[[show_sorts = false]] ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ val thm = @{thm "identity_conc'"}; (*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] []; \ term*\@{thm \Reification_Test.identity_conc\}\ value*\proof @{thm \Reification_Test.identity_conc\}\ lemma cons_list' : "a#xs = [a]@xs" using [[simp_trace]] apply (subst List.append.append_Cons) apply (subst List.append.append_Nil) apply (rule refl) done ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ val thm = @{thm "cons_list'"}; (*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] []; \ declare[[show_sorts = false]] declare[[ML_print_depth = 20]] ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ 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] []; \ prf test full_prf test term*\@{thm \Reification_Test.test\}\ value*\proof @{thm \Reification_Test.test\}\ term*\@{thms-of \Reification_Test.test\}\ value*\@{thms-of \Reification_Test.test\}\ ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ 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] []; \ prf test2 full_prf test2 term*\@{thm \Reification_Test.test2\}\ value*\proof @{thm \Reification_Test.test2\}\ ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ 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] []; \ prf test3 full_prf test3 term*\@{thm \Reification_Test.test3\}\ value*\@{thm \Reification_Test.test3\}\ ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ 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] []; \ prf test4 full_prf test4 term*\@{thm \Reification_Test.test4\}\ value*\@{thm \Reification_Test.test4\}\ ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ 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] []; \ prf symmetric full_prf symmetric term*\@{thm \Pure.symmetric\}\ value*\proof @{thm \Pure.symmetric\}\ ML\ 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 \ ML\ val thm = Proof_Context.get_thm \<^context> "Groups.minus_class.super" val prop = Thm.prop_of thm val proof = Thm.proof_of thm \ prf Groups.minus_class.super full_prf Groups.minus_class.super term*\@{thm \Groups.minus_class.super\}\ value*\@{thm \Groups.minus_class.super\}\ (*ML\ 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 \ ML\ val thm = Proof_Context.get_thm \<^context> "Homotopy.starlike_imp_contractible" val prop = Thm.prop_of thm val proof = Thm.proof_of thm \ prf Homotopy.starlike_imp_contractible full_prf Homotopy.starlike_imp_contractible term*\@{thm \Homotopy.starlike_imp_contractible\}\ value*\@{thm \Homotopy.starlike_imp_contractible\}\*) (* stefan bergofer phd thesis example proof construction 2.3.2 *) lemma stefan_example : "(\x. \y. P x y) \ (\y. \x. P x y)" apply (rule impI) apply(rule allI) apply (erule exE) apply(rule exI) apply(erule allE) apply (assumption) done ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ val thm = @{thm stefan_example}; (*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] []; \ ML\ val thy = \<^theory>; val prf = Proof_Syntax.read_proof thy true false "mp \ _ \ _ \ (impI \ _ \ _ \ (conjI \ _ \ _ ))"; (*"conjI \ _ \ _ ";*) (*"(\<^bold>\(H: _) Ha: _. conjI \ _ \ _ \ Ha \ H)";*) (*val t = Proofterm.reconstruct_proof thy \<^prop>\(A \ B) \ A \ B\ prf*) (* val thm = Proofterm.reconstruct_proof thy \<^prop>\A \ B\ prf |> Proof_Checker.thm_of_proof thy |> Drule.export_without_context val pretty = Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm);*) \ extract_type "typeof (Trueprop P) \ typeof P" realizers impI (P, Q): "\pq. pq" "\<^bold>\(c: _) (d: _) P Q pq (h: _). allI \ _ \ c \ (\<^bold>\x. impI \ _ \ _ \ (h \ x))" find_consts name:"MinProof" ML_val \ val thy = \<^theory>; val prf = Proof_Syntax.read_proof thy true false "impI \ _ \ _ \ \ \ (\<^bold>\H: _. \ \ conjE \ _ \ _ \ _ \ H \ \ \ (\<^bold>\(H: _) Ha: _. conjI \ _ \ _ \ Ha \ H))"; val thm = Proofterm.reconstruct_proof thy \<^prop>\A \ B \ B \ A\ prf |> Proof_Checker.thm_of_proof thy |> Drule.export_without_context; val pretty = Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm); \ ML_file "~~/src/Provers/classical.ML" lemma testtest : "A \ B \ B \ A" apply (rule impI) apply (erule conjE) apply(erule conjI) apply assumption done ML\ (*See: *) \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\\ ML\ val thm = @{thm testtest}; (*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] []; \ ML\ val thy = \<^theory> val prf = Proof_Syntax.read_proof thy true false "impI \ _ \ _ \ \ \ (\<^bold>\H: _. \ \ conjE \ _ \ _ \ _ \ H \ \ \ (\<^bold>\(H: _) Ha: _. conjI \ _ \ _ \ Ha \ H))"; \ ML\ val thy = \<^theory> val prf = Proof_Syntax.read_proof thy true false "\<^bold>\(H: A \ B). conjE \ A \ B \ A \ B \ H"; (* val thm = Proofterm.reconstruct_proof thy \<^prop>\A \ B \ B \ A\ prf |> Proof_Checker.thm_of_proof thy |> Drule.export_without_context; val pretty = Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm);*) \ ML\ val thy = \<^theory> val prf = Proof_Syntax.read_proof thy true false "\<^bold>\(H: _) Ha: _. conjI \ _ \ _ \ Ha \ H"; val thm = Proofterm.reconstruct_proof thy \<^prop>\A \ B \ B \ A\ prf |> Proof_Checker.thm_of_proof thy |> Drule.export_without_context; val pretty = Pretty.writeln (Proof_Syntax.pretty_standard_proof_of \<^context> true thm); \ end