| author | haftmann | 
| Mon, 20 Jul 2009 15:24:15 +0200 | |
| changeset 32082 | 90d03908b3d7 | 
| parent 27034 | 5257bc7e0c06 | 
| child 36862 | 952b2b102a0a | 
| permissions | -rw-r--r-- | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 1 | (* Title: HOL/Induct/PropLog.thy | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 3 | Author: Tobias Nipkow | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 4 | Copyright 1994 TU Muenchen & University of Cambridge | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 5 | *) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 6 | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 7 | header {* Meta-theory of propositional logic *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 8 | |
| 16417 | 9 | theory PropLog imports Main begin | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 10 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 11 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 12 | Datatype definition of propositional logic formulae and inductive | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 13 | definition of the propositional tautologies. | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 14 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 15 | Inductive definition of propositional logic. Soundness and | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 16 | completeness w.r.t.\ truth-tables. | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 17 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 18 |   Prove: If @{text "H |= p"} then @{text "G |= p"} where @{text "G \<in>
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 19 | Fin(H)"} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 20 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 21 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 22 | subsection {* The datatype of propositions *}
 | 
| 5184 | 23 | |
| 24824 | 24 | datatype 'a pl = | 
| 25 | false | | |
| 26 |   var 'a ("#_" [1000]) |
 | |
| 27 | imp "'a pl" "'a pl" (infixr "->" 90) | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 28 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 29 | subsection {* The proof system *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 30 | |
| 23746 | 31 | inductive | 
| 32 | thms :: "['a pl set, 'a pl] => bool" (infixl "|-" 50) | |
| 33 | for H :: "'a pl set" | |
| 34 | where | |
| 35 | H [intro]: "p\<in>H ==> H |- p" | |
| 36 | | K: "H |- p->q->p" | |
| 37 | | S: "H |- (p->q->r) -> (p->q) -> p->r" | |
| 38 | | DN: "H |- ((p->false) -> false) -> p" | |
| 39 | | MP: "[| H |- p->q; H |- p |] ==> H |- q" | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 40 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 41 | subsection {* The semantics *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 42 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 43 | subsubsection {* Semantics of propositional logic. *}
 | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 44 | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 45 | consts | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 46 |   eval  :: "['a set, 'a pl] => bool"      ("_[[_]]" [100,0] 100)
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 47 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 48 | primrec "tt[[false]] = False" | 
| 18260 | 49 | "tt[[#v]] = (v \<in> tt)" | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 50 | eval_imp: "tt[[p->q]] = (tt[[p]] --> tt[[q]])" | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 51 | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 52 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 53 |   A finite set of hypotheses from @{text t} and the @{text Var}s in
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 54 |   @{text p}.
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 55 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 56 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 57 | consts | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 58 | hyps :: "['a pl, 'a set] => 'a pl set" | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 59 | |
| 5184 | 60 | primrec | 
| 10759 | 61 |   "hyps false  tt = {}"
 | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 62 |   "hyps (#v)   tt = {if v \<in> tt then #v else #v->false}"
 | 
| 10759 | 63 | "hyps (p->q) tt = hyps p tt Un hyps q tt" | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 64 | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 65 | subsubsection {* Logical consequence *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 66 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 67 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 68 |   For every valuation, if all elements of @{text H} are true then so
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 69 |   is @{text p}.
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 70 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 71 | |
| 19736 | 72 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20503diff
changeset | 73 | sat :: "['a pl set, 'a pl] => bool" (infixl "|=" 50) where | 
| 19736 | 74 | "H |= p = (\<forall>tt. (\<forall>q\<in>H. tt[[q]]) --> tt[[p]])" | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 75 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 76 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 77 | subsection {* Proof theory of propositional logic *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 78 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 79 | lemma thms_mono: "G<=H ==> thms(G) <= thms(H)" | 
| 23746 | 80 | apply (rule predicate1I) | 
| 81 | apply (erule thms.induct) | |
| 82 | apply (auto intro: thms.intros) | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 83 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 84 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 85 | lemma thms_I: "H |- p->p" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 86 |   -- {*Called @{text I} for Identity Combinator, not for Introduction. *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 87 | by (best intro: thms.K thms.S thms.MP) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 88 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 89 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 90 | subsubsection {* Weakening, left and right *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 91 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 92 | lemma weaken_left: "[| G \<subseteq> H; G|-p |] ==> H|-p" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 93 |   -- {* Order of premises is convenient with @{text THEN} *}
 | 
| 23746 | 94 | by (erule thms_mono [THEN predicate1D]) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 95 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 96 | lemmas weaken_left_insert = subset_insertI [THEN weaken_left] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 97 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 98 | lemmas weaken_left_Un1 = Un_upper1 [THEN weaken_left] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 99 | lemmas weaken_left_Un2 = Un_upper2 [THEN weaken_left] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 100 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 101 | lemma weaken_right: "H |- q ==> H |- p->q" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 102 | by (fast intro: thms.K thms.MP) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 103 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 104 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 105 | subsubsection {* The deduction theorem *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 106 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 107 | theorem deduction: "insert p H |- q ==> H |- p->q" | 
| 18260 | 108 | apply (induct set: thms) | 
| 109 | apply (fast intro: thms_I thms.H thms.K thms.S thms.DN | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 110 | thms.S [THEN thms.MP, THEN thms.MP] weaken_right)+ | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 111 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 112 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 113 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 114 | subsubsection {* The cut rule *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 115 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 116 | lemmas cut = deduction [THEN thms.MP] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 117 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 118 | lemmas thms_falseE = weaken_right [THEN thms.DN [THEN thms.MP]] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 119 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 120 | lemmas thms_notE = thms.MP [THEN thms_falseE, standard] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 121 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 122 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 123 | subsubsection {* Soundness of the rules wrt truth-table semantics *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 124 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 125 | theorem soundness: "H |- p ==> H |= p" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 126 | apply (unfold sat_def) | 
| 18260 | 127 | apply (induct set: thms) | 
| 128 | apply auto | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 129 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 130 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 131 | subsection {* Completeness *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 132 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 133 | subsubsection {* Towards the completeness proof *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 134 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 135 | lemma false_imp: "H |- p->false ==> H |- p->q" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 136 | apply (rule deduction) | 
| 27034 | 137 | apply (metis H insert_iff weaken_left_insert thms_notE) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 138 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 139 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 140 | lemma imp_false: | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 141 | "[| H |- p; H |- q->false |] ==> H |- (p->q)->false" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 142 | apply (rule deduction) | 
| 27034 | 143 | apply (metis H MP insert_iff weaken_left_insert) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 144 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 145 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 146 | lemma hyps_thms_if: "hyps p tt |- (if tt[[p]] then p else p->false)" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 147 |   -- {* Typical example of strengthening the induction statement. *}
 | 
| 18260 | 148 | apply simp | 
| 149 | apply (induct p) | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 150 | apply (simp_all add: thms_I thms.H) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 151 | apply (blast intro: weaken_left_Un1 weaken_left_Un2 weaken_right | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 152 | imp_false false_imp) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 153 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 154 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 155 | lemma sat_thms_p: "{} |= p ==> hyps p tt |- p"
 | 
| 18260 | 156 |   -- {* Key lemma for completeness; yields a set of assumptions
 | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 157 |         satisfying @{text p} *}
 | 
| 18260 | 158 | apply (unfold sat_def) | 
| 159 | apply (drule spec, erule mp [THEN if_P, THEN subst], | |
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 160 | rule_tac [2] hyps_thms_if, simp) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 161 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 162 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 163 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 164 | For proving certain theorems in our new propositional logic. | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 165 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 166 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 167 | declare deduction [intro!] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 168 | declare thms.H [THEN thms.MP, intro] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 169 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 170 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 171 | The excluded middle in the form of an elimination rule. | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 172 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 173 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 174 | lemma thms_excluded_middle: "H |- (p->q) -> ((p->false)->q) -> q" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 175 | apply (rule deduction [THEN deduction]) | 
| 18260 | 176 | apply (rule thms.DN [THEN thms.MP], best) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 177 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 178 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 179 | lemma thms_excluded_middle_rule: | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 180 | "[| insert p H |- q; insert (p->false) H |- q |] ==> H |- q" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 181 |   -- {* Hard to prove directly because it requires cuts *}
 | 
| 18260 | 182 | by (rule thms_excluded_middle [THEN thms.MP, THEN thms.MP], auto) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 183 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 184 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 185 | subsection{* Completeness -- lemmas for reducing the set of assumptions*}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 186 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 187 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 188 |   For the case @{prop "hyps p t - insert #v Y |- p"} we also have @{prop
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 189 |   "hyps p t - {#v} \<subseteq> hyps p (t-{v})"}.
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 190 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 191 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 192 | lemma hyps_Diff: "hyps p (t-{v}) <= insert (#v->false) ((hyps p t)-{#v})"
 | 
| 18260 | 193 | by (induct p) auto | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 194 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 195 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 196 |   For the case @{prop "hyps p t - insert (#v -> Fls) Y |- p"} we also have
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 197 |   @{prop "hyps p t-{#v->Fls} \<subseteq> hyps p (insert v t)"}.
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 198 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 199 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 200 | lemma hyps_insert: "hyps p (insert v t) <= insert (#v) (hyps p t-{#v->false})"
 | 
| 18260 | 201 | by (induct p) auto | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 202 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 203 | text {* Two lemmas for use with @{text weaken_left} *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 204 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 205 | lemma insert_Diff_same: "B-C <= insert a (B-insert a C)" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 206 | by fast | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 207 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 208 | lemma insert_Diff_subset2: "insert a (B-{c}) - D <= insert a (B-insert c D)"
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 209 | by fast | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 210 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 211 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 212 |   The set @{term "hyps p t"} is finite, and elements have the form
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 213 |   @{term "#v"} or @{term "#v->Fls"}.
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 214 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 215 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 216 | lemma hyps_finite: "finite(hyps p t)" | 
| 18260 | 217 | by (induct p) auto | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 218 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 219 | lemma hyps_subset: "hyps p t <= (UN v. {#v, #v->false})"
 | 
| 18260 | 220 | by (induct p) auto | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 221 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 222 | lemmas Diff_weaken_left = Diff_mono [OF _ subset_refl, THEN weaken_left] | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 223 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 224 | subsubsection {* Completeness theorem *}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 225 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 226 | text {*
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 227 |   Induction on the finite set of assumptions @{term "hyps p t0"}.  We
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 228 | may repeatedly subtract assumptions until none are left! | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 229 | *} | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 230 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 231 | lemma completeness_0_lemma: | 
| 18260 | 232 |     "{} |= p ==> \<forall>t. hyps p t - hyps p t0 |- p"
 | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 233 | apply (rule hyps_subset [THEN hyps_finite [THEN finite_subset_induct]]) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 234 | apply (simp add: sat_thms_p, safe) | 
| 16563 | 235 |  txt{*Case @{text"hyps p t-insert(#v,Y) |- p"} *}
 | 
| 18260 | 236 | apply (iprover intro: thms_excluded_middle_rule | 
| 237 | insert_Diff_same [THEN weaken_left] | |
| 238 | insert_Diff_subset2 [THEN weaken_left] | |
| 16563 | 239 | hyps_Diff [THEN Diff_weaken_left]) | 
| 240 | txt{*Case @{text"hyps p t-insert(#v -> false,Y) |- p"} *}
 | |
| 18260 | 241 | apply (iprover intro: thms_excluded_middle_rule | 
| 242 | insert_Diff_same [THEN weaken_left] | |
| 243 | insert_Diff_subset2 [THEN weaken_left] | |
| 16563 | 244 | hyps_insert [THEN Diff_weaken_left]) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 245 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 246 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 247 | text{*The base case for completeness*}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 248 | lemma completeness_0:  "{} |= p ==> {} |- p"
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 249 | apply (rule Diff_cancel [THEN subst]) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 250 | apply (erule completeness_0_lemma [THEN spec]) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 251 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 252 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 253 | text{*A semantic analogue of the Deduction Theorem*}
 | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 254 | lemma sat_imp: "insert p H |= q ==> H |= p->q" | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 255 | by (unfold sat_def, auto) | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 256 | |
| 18260 | 257 | theorem completeness: "finite H ==> H |= p ==> H |- p" | 
| 20503 | 258 | apply (induct arbitrary: p rule: finite_induct) | 
| 16563 | 259 | apply (blast intro: completeness_0) | 
| 17589 | 260 | apply (iprover intro: sat_imp thms.H insertI1 weaken_left_insert [THEN thms.MP]) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 261 | done | 
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 262 | |
| 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 263 | theorem syntax_iff_semantics: "finite H ==> (H |- p) = (H |= p)" | 
| 16563 | 264 | by (blast intro: soundness completeness) | 
| 13075 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 paulson parents: 
10759diff
changeset | 265 | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 266 | end |