| author | hoelzl | 
| Thu, 26 May 2011 14:12:03 +0200 | |
| changeset 42986 | 11fd8c04ea24 | 
| parent 42616 | 92715b528e78 | 
| child 43016 | 42330f25142c | 
| 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 | ||
| 7 | header {* Metis Proof Method *}
 | |
| 8 | ||
| 9 | theory Metis | |
| 10 | imports Meson | |
| 11 | uses "~~/src/Tools/Metis/metis.ML" | |
| 12 |      ("Tools/Metis/metis_translate.ML")
 | |
| 13 |      ("Tools/Metis/metis_reconstruct.ML")
 | |
| 14 |      ("Tools/Metis/metis_tactics.ML")
 | |
| 41042 
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
 blanchet parents: 
39980diff
changeset | 15 |      ("Tools/try.ML")
 | 
| 39946 | 16 | begin | 
| 17 | ||
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 18 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 19 | subsection {* Higher-order reasoning helpers *}
 | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 20 | |
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 21 | definition fFalse :: bool where [no_atp]: | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 22 | "fFalse \<longleftrightarrow> False" | 
| 39946 | 23 | |
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 24 | definition fTrue :: bool where [no_atp]: | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 25 | "fTrue \<longleftrightarrow> True" | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 26 | |
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 27 | definition fNot :: "bool \<Rightarrow> bool" where [no_atp]: | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 28 | "fNot P \<longleftrightarrow> \<not> P" | 
| 39946 | 29 | |
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 30 | definition fconj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]: | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 31 | "fconj P Q \<longleftrightarrow> P \<and> Q" | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 32 | |
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 33 | definition fdisj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]: | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 34 | "fdisj P Q \<longleftrightarrow> P \<or> Q" | 
| 39946 | 35 | |
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 36 | definition fimplies :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]: | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 37 | "fimplies P Q \<longleftrightarrow> (P \<longrightarrow> Q)" | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 38 | |
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 39 | definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where [no_atp]: | 
| 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 40 | "fequal x y \<longleftrightarrow> (x = y)" | 
| 39946 | 41 | |
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 42 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 43 | subsection {* Literal selection helpers *}
 | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 44 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 45 | definition select :: "'a \<Rightarrow> 'a" where | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 46 | [no_atp]: "select = (\<lambda>x. x)" | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 47 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 48 | lemma not_atomize: "(\<not> A \<Longrightarrow> False) \<equiv> Trueprop A" | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 49 | by (cut_tac atomize_not [of "\<not> A"]) simp | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 50 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 51 | 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 | 52 | unfolding select_def by (rule atomize_not) | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 53 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 54 | 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 | 55 | unfolding select_def by (rule not_atomize) | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 56 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 57 | lemma select_FalseI: "False \<Longrightarrow> select False" by simp | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 58 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 59 | |
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 60 | subsection {* Metis package *}
 | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 61 | |
| 39946 | 62 | use "Tools/Metis/metis_translate.ML" | 
| 63 | use "Tools/Metis/metis_reconstruct.ML" | |
| 64 | use "Tools/Metis/metis_tactics.ML" | |
| 39980 | 65 | |
| 42616 
92715b528e78
added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
 wenzelm parents: 
42349diff
changeset | 66 | setup {* Metis_Tactics.setup *}
 | 
| 39946 | 67 | |
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 68 | hide_const (open) fFalse fTrue fNot fconj fdisj fimplies fequal select | 
| 41140 
9c68004b8c9d
added Sledgehammer support for higher-order propositional reasoning
 blanchet parents: 
41042diff
changeset | 69 | hide_fact (open) fFalse_def fTrue_def fNot_def fconj_def fdisj_def fimplies_def | 
| 42349 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 70 | fequal_def select_def not_atomize atomize_not_select not_atomize_select | 
| 
721e85fd2db3
make 48170228f562 work also with "HO_Reas" examples
 blanchet parents: 
41140diff
changeset | 71 | select_FalseI | 
| 39953 
aa54f347e5e2
hide uninteresting MESON/Metis constants and facts and remove "meson_" prefix to (now hidden) fact names
 blanchet parents: 
39947diff
changeset | 72 | |
| 41042 
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
 blanchet parents: 
39980diff
changeset | 73 | subsection {* Try *}
 | 
| 
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
 blanchet parents: 
39980diff
changeset | 74 | |
| 
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
 blanchet parents: 
39980diff
changeset | 75 | use "Tools/try.ML" | 
| 
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
 blanchet parents: 
39980diff
changeset | 76 | |
| 
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
 blanchet parents: 
39980diff
changeset | 77 | setup {* Try.setup *}
 | 
| 
8275f52ac991
load "try" after "Metis" and move "Async_Manager" back to Sledgehammer
 blanchet parents: 
39980diff
changeset | 78 | |
| 39946 | 79 | end |