| author | wenzelm | 
| Fri, 08 Dec 2023 15:37:46 +0100 | |
| changeset 79207 | f991d3003ec8 | 
| parent 78728 | 72631efa3821 | 
| child 81254 | d3c0734059ee | 
| permissions | -rw-r--r-- | 
| 39946 | 1 | (* Title: HOL/Metis.thy | 
| 2 | Author: Lawrence C. Paulson, Cambridge University Computer Laboratory | |
| 3 | Author: Jia Meng, Cambridge University Computer Laboratory and NICTA | |
| 4 | Author: Jasmin Blanchette, TU Muenchen | |
| 5 | *) | |
| 6 | ||
| 60758 | 7 | section \<open>Metis Proof Method\<close> | 
| 39946 | 8 | |
| 9 | theory Metis | |
| 43085 
0a2f5b86bdd7
first step in sharing more code between ATP and Metis translation
 blanchet parents: 
43016diff
changeset | 10 | imports ATP | 
| 39946 | 11 | begin | 
| 12 | ||
| 78728 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
77263diff
changeset | 13 | context notes [[ML_catch_all]] | 
| 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
77263diff
changeset | 14 | begin | 
| 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
77263diff
changeset | 15 | ML_file \<open>~~/src/Tools/Metis/metis.ML\<close> | 
| 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
77263diff
changeset | 16 | end | 
| 48891 | 17 | |
| 56946 | 18 | |
| 60758 | 19 | subsection \<open>Literal selection and lambda-lifting helpers\<close> | 
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 20 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 21 | definition select :: "'a \<Rightarrow> 'a" where | 
| 54148 | 22 | "select = (\<lambda>x. x)" | 
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 23 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 24 | lemma not_atomize: "(\<not> A \<Longrightarrow> False) \<equiv> Trueprop A" | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 25 | by (cut_tac atomize_not [of "\<not> A"]) simp | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 26 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 27 | lemma atomize_not_select: "(A \<Longrightarrow> select False) \<equiv> Trueprop (\<not> A)" | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 28 | unfolding select_def by (rule atomize_not) | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 29 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 30 | lemma not_atomize_select: "(\<not> A \<Longrightarrow> select False) \<equiv> Trueprop A" | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 31 | unfolding select_def by (rule not_atomize) | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 32 | |
| 77263 
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
 blanchet parents: 
69605diff
changeset | 33 | lemma select_FalseI: "False \<Longrightarrow> select False" | 
| 
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
 blanchet parents: 
69605diff
changeset | 34 | by simp | 
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 35 | |
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
44651diff
changeset | 36 | definition lambda :: "'a \<Rightarrow> 'a" where | 
| 54148 | 37 | "lambda = (\<lambda>x. x)" | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
44651diff
changeset | 38 | |
| 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
44651diff
changeset | 39 | lemma eq_lambdaI: "x \<equiv> y \<Longrightarrow> x \<equiv> lambda y" | 
| 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
44651diff
changeset | 40 | unfolding lambda_def by assumption | 
| 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
44651diff
changeset | 41 | |
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 42 | |
| 60758 | 43 | subsection \<open>Metis package\<close> | 
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 44 | |
| 69605 | 45 | ML_file \<open>Tools/Metis/metis_generate.ML\<close> | 
| 46 | ML_file \<open>Tools/Metis/metis_reconstruct.ML\<close> | |
| 47 | ML_file \<open>Tools/Metis/metis_tactic.ML\<close> | |
| 39980 | 48 | |
| 56946 | 49 | hide_const (open) select fFalse fTrue fNot fComp fconj fdisj fimplies fAll fEx fequal lambda | 
| 50 | hide_fact (open) select_def not_atomize atomize_not_select not_atomize_select select_FalseI | |
| 51 | fFalse_def fTrue_def fNot_def fconj_def fdisj_def fimplies_def fAll_def fEx_def fequal_def | |
| 52 | fTrue_ne_fFalse fNot_table fconj_table fdisj_table fimplies_table fAll_table fEx_table | |
| 53 | fequal_table fAll_table fEx_table fNot_law fComp_law fconj_laws fdisj_laws fimplies_laws | |
| 54 | fequal_laws fAll_law fEx_law lambda_def eq_lambdaI | |
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
44651diff
changeset | 55 | |
| 39946 | 56 | end |