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