| author | wenzelm | 
| Thu, 08 Jul 2021 22:21:31 +0200 | |
| changeset 73950 | cc49da3003aa | 
| parent 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 62287 | 1 | (* Title: HOL/Eisbach/Examples_FOL.thy | 
| 62168 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 2 | Author: Daniel Matichuk, NICTA/UNSW | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 3 | *) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 4 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 5 | section \<open>Basic Eisbach examples (in FOL)\<close> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 6 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 7 | theory Examples_FOL | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
62287diff
changeset | 8 | imports FOL Eisbach_Old_Appl_Syntax | 
| 62168 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 9 | begin | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 10 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 11 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 12 | subsection \<open>Basic methods\<close> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 13 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 14 | method my_intros = (rule conjI | rule impI) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 15 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 16 | lemma "P \<and> Q \<longrightarrow> Z \<and> X" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 17 | apply my_intros+ | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 18 | oops | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 19 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 20 | method my_intros' uses intros = (rule conjI | rule impI | rule intros) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 21 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 22 | lemma "P \<and> Q \<longrightarrow> Z \<or> X" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 23 | apply (my_intros' intros: disjI1)+ | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 24 | oops | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 25 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 26 | method my_spec for x :: 'a = (drule spec[where x="x"]) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 27 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 28 | lemma "\<forall>x. P(x) \<Longrightarrow> P(x)" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 29 | apply (my_spec x) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 30 | apply assumption | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 31 | done | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 32 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 33 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 34 | subsection \<open>Demo\<close> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 35 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 36 | named_theorems intros and elims and subst | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 37 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 38 | method prop_solver declares intros elims subst = | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 39 | (assumption | | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 40 | rule intros | erule elims | | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 41 | subst subst | subst (asm) subst | | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 42 | (erule notE; solves prop_solver))+ | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 43 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 44 | lemmas [intros] = | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 45 | conjI | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 46 | impI | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 47 | disjCI | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 48 | iffI | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 49 | notI | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 50 | lemmas [elims] = | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 51 | impCE | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 52 | conjE | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 53 | disjE | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 54 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 55 | lemma "((A \<or> B) \<and> (A \<longrightarrow> C) \<and> (B \<longrightarrow> C)) \<longrightarrow> C" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 56 | apply prop_solver | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 57 | done | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 58 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 59 | method guess_all = | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 60 | (match premises in U[thin]:"\<forall>x. P (x :: 'a)" for P \<Rightarrow> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 61 | \<open>match premises in "?H (y :: 'a)" for y \<Rightarrow> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 62 | \<open>rule allE[where P = P and x = y, OF U]\<close> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 63 | | match conclusion in "?H (y :: 'a)" for y \<Rightarrow> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 64 | \<open>rule allE[where P = P and x = y, OF U]\<close>\<close>) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 65 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 66 | lemma "(\<forall>x. P(x) \<longrightarrow> Q(x)) \<Longrightarrow> P(y) \<Longrightarrow> Q(y)" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 67 | apply guess_all | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 68 | apply prop_solver | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 69 | done | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 70 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 71 | lemma "(\<forall>x. P(x) \<longrightarrow> Q(x)) \<Longrightarrow> P(z) \<Longrightarrow> P(y) \<Longrightarrow> Q(y)" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 72 | apply (solves \<open>guess_all, prop_solver\<close>) \<comment> \<open>Try it without solve\<close> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 73 | done | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 74 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 75 | method guess_ex = | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 76 | (match conclusion in | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 77 | "\<exists>x. P (x :: 'a)" for P \<Rightarrow> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 78 | \<open>match premises in "?H (x :: 'a)" for x \<Rightarrow> | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 79 | \<open>rule exI[where x=x]\<close>\<close>) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 80 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 81 | lemma "P(x) \<Longrightarrow> \<exists>x. P(x)" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 82 | apply guess_ex | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 83 | apply prop_solver | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 84 | done | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 85 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 86 | method fol_solver = | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 87 | ((guess_ex | guess_all | prop_solver); solves fol_solver) | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 88 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 89 | declare | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 90 | allI [intros] | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 91 | exE [elims] | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 92 | ex_simps [subst] | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 93 | all_simps [subst] | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 94 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 95 | lemma "(\<forall>x. P(x)) \<and> (\<forall>x. Q(x)) \<Longrightarrow> (\<forall>x. P(x) \<and> Q(x))" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 96 | and "\<exists>x. P(x) \<longrightarrow> (\<forall>x. P(x))" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 97 | and "(\<exists>x. \<forall>y. R(x, y)) \<longrightarrow> (\<forall>y. \<exists>x. R(x, y))" | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 98 | by fol_solver+ | 
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 99 | |
| 
e97452d79102
Eisbach works for other object-logics, e.g. Eisbach_FOL.thy;
 wenzelm parents: diff
changeset | 100 | end |