| author | nipkow | 
| Tue, 14 Feb 2023 08:10:17 +0100 | |
| changeset 77266 | 334015f9098e | 
| parent 77062 | 1d5872cb52ec | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
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: 
10759 
diff
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: 
10759 
diff
changeset
 | 
7  | 
|
| 16417 | 8  | 
theory PropLog imports Main begin  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
9  | 
|
| 60530 | 10  | 
text \<open>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
11  | 
Datatype definition of propositional logic formulae and inductive  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
12  | 
definition of the propositional tautologies.  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
13  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
14  | 
Inductive definition of propositional logic. Soundness and  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
15  | 
completeness w.r.t.\ truth-tables.  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
16  | 
|
| 77062 | 17  | 
Prove: If \<open>H \<Turnstile> p\<close> then \<open>G \<Turnstile> p\<close> where \<open>G \<in>  | 
| 63167 | 18  | 
Fin(H)\<close>  | 
| 60530 | 19  | 
\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
20  | 
|
| 60530 | 21  | 
subsection \<open>The datatype of propositions\<close>  | 
| 5184 | 22  | 
|
| 58310 | 23  | 
datatype 'a pl =  | 
| 77062 | 24  | 
false  | 
25  | 
  | var 'a ("#_" [1000]) 
 | 
|
26  | 
| imp "'a pl" "'a pl" (infixr "\<rightharpoonup>" 90)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
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: 
10759 
diff
changeset
 | 
30  | 
|
| 77062 | 31  | 
inductive thms :: "['a pl set, 'a pl] \<Rightarrow> bool" (infixl "\<turnstile>" 50)  | 
| 23746 | 32  | 
for H :: "'a pl set"  | 
| 77062 | 33  | 
where  | 
34  | 
H: "p \<in> H \<Longrightarrow> H \<turnstile> p"  | 
|
35  | 
| K: "H \<turnstile> p\<rightharpoonup>q\<rightharpoonup>p"  | 
|
36  | 
| S: "H \<turnstile> (p\<rightharpoonup>q\<rightharpoonup>r) \<rightharpoonup> (p\<rightharpoonup>q) \<rightharpoonup> p\<rightharpoonup>r"  | 
|
37  | 
| DN: "H \<turnstile> ((p\<rightharpoonup>false) \<rightharpoonup> false) \<rightharpoonup> p"  | 
|
38  | 
| MP: "\<lbrakk>H \<turnstile> p\<rightharpoonup>q; H \<turnstile> p\<rbrakk> \<Longrightarrow> H \<turnstile> q"  | 
|
| 46579 | 39  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
40  | 
|
| 60530 | 41  | 
subsection \<open>The semantics\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
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)
 | 
| 77062 | 46  | 
where  | 
47  | 
"tt[[false]] = False"  | 
|
48  | 
| "tt[[#v]] = (v \<in> tt)"  | 
|
49  | 
| eval_imp: "tt[[p\<rightharpoonup>q]] = (tt[[p]] \<longrightarrow> 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: 
10759 
diff
changeset
 | 
55  | 
|
| 46579 | 56  | 
primrec hyps :: "['a pl, 'a set] => 'a pl set"  | 
| 77062 | 57  | 
where  | 
58  | 
    "hyps false  tt = {}"
 | 
|
59  | 
  | "hyps (#v)   tt = {if v \<in> tt then #v else #v\<rightharpoonup>false}"
 | 
|
60  | 
| "hyps (p\<rightharpoonup>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: 
10759 
diff
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: 
10759 
diff
changeset
 | 
69  | 
|
| 77062 | 70  | 
definition sat :: "['a pl set, 'a pl] => bool" (infixl "\<Turnstile>" 50)  | 
71  | 
where "H \<Turnstile> p = (\<forall>tt. (\<forall>q\<in>H. tt[[q]]) \<longrightarrow> tt[[p]])"  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
72  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
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: 
10759 
diff
changeset
 | 
75  | 
|
| 77062 | 76  | 
lemma thms_mono:  | 
77  | 
assumes "G \<subseteq> H" shows "thms(G) \<le> thms(H)"  | 
|
78  | 
proof -  | 
|
79  | 
have "G \<turnstile> p \<Longrightarrow> H \<turnstile> p" for p  | 
|
80  | 
by (induction rule: thms.induct) (use assms in \<open>auto intro: thms.intros\<close>)  | 
|
81  | 
then show ?thesis  | 
|
82  | 
by blast  | 
|
83  | 
qed  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
84  | 
|
| 77062 | 85  | 
lemma thms_I: "H \<turnstile> p\<rightharpoonup>p"  | 
| 63167 | 86  | 
\<comment> \<open>Called \<open>I\<close> for Identity Combinator, not for Introduction.\<close>  | 
| 77062 | 87  | 
by (best intro: thms.K thms.S thms.MP)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
88  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
89  | 
|
| 60530 | 90  | 
subsubsection \<open>Weakening, left and right\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
91  | 
|
| 77062 | 92  | 
lemma weaken_left: "\<lbrakk>G \<subseteq> H; G\<turnstile>p\<rbrakk> \<Longrightarrow> H\<turnstile>p"  | 
| 63167 | 93  | 
\<comment> \<open>Order of premises is convenient with \<open>THEN\<close>\<close>  | 
| 77062 | 94  | 
by (meson predicate1D thms_mono)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
95  | 
|
| 77062 | 96  | 
lemma weaken_left_insert: "G \<turnstile> p \<Longrightarrow> insert a G \<turnstile> p"  | 
97  | 
by (meson subset_insertI weaken_left)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
98  | 
|
| 77062 | 99  | 
lemma weaken_left_Un1: "G \<turnstile> p \<Longrightarrow> G \<union> B \<turnstile> p"  | 
100  | 
by (rule weaken_left) (rule Un_upper1)  | 
|
| 46579 | 101  | 
|
| 77062 | 102  | 
lemma weaken_left_Un2: "G \<turnstile> p \<Longrightarrow> A \<union> G \<turnstile> p"  | 
103  | 
by (metis Un_commute weaken_left_Un1)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
104  | 
|
| 77062 | 105  | 
lemma weaken_right: "H \<turnstile> q \<Longrightarrow> H \<turnstile> p\<rightharpoonup>q"  | 
106  | 
using K MP by blast  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
107  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
108  | 
|
| 60530 | 109  | 
subsubsection \<open>The deduction theorem\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
110  | 
|
| 77062 | 111  | 
theorem deduction: "insert p H \<turnstile> q \<Longrightarrow> H \<turnstile> p\<rightharpoonup>q"  | 
112  | 
proof (induct set: thms)  | 
|
113  | 
case (H p)  | 
|
114  | 
then show ?case  | 
|
115  | 
using thms.H thms_I weaken_right by fastforce  | 
|
116  | 
qed (metis thms.simps)+  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
117  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
118  | 
|
| 60530 | 119  | 
subsubsection \<open>The cut rule\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
120  | 
|
| 77062 | 121  | 
lemma cut: "insert p H \<turnstile> q \<Longrightarrow> H \<turnstile> p \<Longrightarrow> H \<turnstile> q"  | 
122  | 
using MP deduction by blast  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
123  | 
|
| 77062 | 124  | 
lemma thms_falseE: "H \<turnstile> false \<Longrightarrow> H \<turnstile> q"  | 
125  | 
by (metis thms.simps)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
126  | 
|
| 77062 | 127  | 
lemma thms_notE: "H \<turnstile> p \<rightharpoonup> false \<Longrightarrow> H \<turnstile> p \<Longrightarrow> H \<turnstile> q"  | 
128  | 
using MP thms_falseE by blast  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
129  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
130  | 
|
| 60530 | 131  | 
subsubsection \<open>Soundness of the rules wrt truth-table semantics\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
132  | 
|
| 77062 | 133  | 
theorem soundness: "H \<turnstile> p \<Longrightarrow> H \<Turnstile> p"  | 
134  | 
by (induct set: thms) (auto simp: sat_def)  | 
|
| 46579 | 135  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
136  | 
|
| 60530 | 137  | 
subsection \<open>Completeness\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
138  | 
|
| 60530 | 139  | 
subsubsection \<open>Towards the completeness proof\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
140  | 
|
| 77062 | 141  | 
lemma false_imp: "H \<turnstile> p\<rightharpoonup>false \<Longrightarrow> H \<turnstile> p\<rightharpoonup>q"  | 
142  | 
by (metis thms.simps)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
143  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
144  | 
lemma imp_false:  | 
| 77062 | 145  | 
"\<lbrakk>H \<turnstile> p; H \<turnstile> q\<rightharpoonup>false\<rbrakk> \<Longrightarrow> H \<turnstile> (p\<rightharpoonup>q)\<rightharpoonup>false"  | 
146  | 
by (meson MP S weaken_right)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
147  | 
|
| 77062 | 148  | 
lemma hyps_thms_if: "hyps p tt \<turnstile> (if tt[[p]] then p else p\<rightharpoonup>false)"  | 
| 63167 | 149  | 
\<comment> \<open>Typical example of strengthening the induction statement.\<close>  | 
| 77062 | 150  | 
proof (induction p)  | 
151  | 
case (imp p1 p2)  | 
|
152  | 
then show ?case  | 
|
153  | 
by (metis (full_types) eval_imp false_imp hyps.simps(3) imp_false weaken_left_Un1 weaken_left_Un2 weaken_right)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
154  | 
|
| 77062 | 155  | 
qed (simp_all add: thms_I thms.H)  | 
156  | 
||
157  | 
lemma sat_thms_p: "{} \<Turnstile> p \<Longrightarrow> hyps p tt \<turnstile> p"
 | 
|
| 63167 | 158  | 
\<comment> \<open>Key lemma for completeness; yields a set of assumptions  | 
159  | 
satisfying \<open>p\<close>\<close>  | 
|
| 77062 | 160  | 
by (metis (full_types) empty_iff hyps_thms_if sat_def)  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
161  | 
|
| 60530 | 162  | 
text \<open>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
163  | 
For proving certain theorems in our new propositional logic.  | 
| 60530 | 164  | 
\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
165  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
166  | 
declare deduction [intro!]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
167  | 
declare thms.H [THEN thms.MP, intro]  | 
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
168  | 
|
| 60530 | 169  | 
text \<open>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
170  | 
The excluded middle in the form of an elimination rule.  | 
| 60530 | 171  | 
\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
172  | 
|
| 77062 | 173  | 
lemma thms_excluded_middle: "H \<turnstile> (p\<rightharpoonup>q) \<rightharpoonup> ((p\<rightharpoonup>false)\<rightharpoonup>q) \<rightharpoonup> q"  | 
174  | 
proof -  | 
|
175  | 
have "insert ((p \<rightharpoonup> false) \<rightharpoonup> q) (insert (p \<rightharpoonup> q) H) \<turnstile> (q \<rightharpoonup> false) \<rightharpoonup> false"  | 
|
176  | 
by (best intro: H)  | 
|
177  | 
then show ?thesis  | 
|
178  | 
by (metis deduction thms.simps)  | 
|
179  | 
qed  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
180  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
181  | 
lemma thms_excluded_middle_rule:  | 
| 77062 | 182  | 
"\<lbrakk>insert p H \<turnstile> q; insert (p\<rightharpoonup>false) H \<turnstile> q\<rbrakk> \<Longrightarrow> H \<turnstile> q"  | 
| 63167 | 183  | 
\<comment> \<open>Hard to prove directly because it requires cuts\<close>  | 
| 77062 | 184  | 
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: 
10759 
diff
changeset
 | 
185  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
186  | 
|
| 60530 | 187  | 
subsection\<open>Completeness -- lemmas for reducing the set of assumptions\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
188  | 
|
| 60530 | 189  | 
text \<open>  | 
| 77062 | 190  | 
  For the case \<^prop>\<open>hyps p t - insert #v Y \<turnstile> p\<close> we also have \<^prop>\<open>hyps p t - {#v} \<subseteq> hyps p (t-{v})\<close>.
 | 
| 60530 | 191  | 
\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
192  | 
|
| 77062 | 193  | 
lemma hyps_Diff: "hyps p (t-{v}) \<subseteq> insert (#v\<rightharpoonup>false) ((hyps p t)-{#v})"
 | 
194  | 
by (induct p) auto  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
195  | 
|
| 60530 | 196  | 
text \<open>  | 
| 77062 | 197  | 
For the case \<^prop>\<open>hyps p t - insert (#v \<rightharpoonup> Fls) Y \<turnstile> p\<close> we also have  | 
198  | 
  \<^prop>\<open>hyps p t-{#v\<rightharpoonup>Fls} \<subseteq> hyps p (insert v t)\<close>.
 | 
|
| 60530 | 199  | 
\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
200  | 
|
| 77062 | 201  | 
lemma hyps_insert: "hyps p (insert v t) \<subseteq> insert (#v) (hyps p t-{#v\<rightharpoonup>false})"
 | 
202  | 
by (induct p) auto  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
203  | 
|
| 63167 | 204  | 
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: 
10759 
diff
changeset
 | 
205  | 
|
| 77062 | 206  | 
lemma insert_Diff_same: "B-C \<subseteq> insert a (B-insert a C)"  | 
207  | 
by fast  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
208  | 
|
| 77062 | 209  | 
lemma insert_Diff_subset2: "insert a (B-{c}) - D \<subseteq> insert a (B-insert c D)"
 | 
210  | 
by fast  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
211  | 
|
| 60530 | 212  | 
text \<open>  | 
| 69597 | 213  | 
The set \<^term>\<open>hyps p t\<close> is finite, and elements have the form  | 
| 77062 | 214  | 
\<^term>\<open>#v\<close> or \<^term>\<open>#v\<rightharpoonup>Fls\<close>.  | 
| 60530 | 215  | 
\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
216  | 
|
| 
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
217  | 
lemma hyps_finite: "finite(hyps p t)"  | 
| 77062 | 218  | 
by (induct p) auto  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
219  | 
|
| 77062 | 220  | 
lemma hyps_subset: "hyps p t \<subseteq> (UN v. {#v, #v\<rightharpoonup>false})"
 | 
221  | 
by (induct p) auto  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
222  | 
|
| 77062 | 223  | 
lemma Diff_weaken_left: "A \<subseteq> C \<Longrightarrow> A - B \<turnstile> p \<Longrightarrow> C - B \<turnstile> p"  | 
| 46579 | 224  | 
by (rule Diff_mono [OF _ subset_refl, THEN weaken_left])  | 
225  | 
||
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
226  | 
|
| 60530 | 227  | 
subsubsection \<open>Completeness theorem\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
228  | 
|
| 60530 | 229  | 
text \<open>  | 
| 69597 | 230  | 
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: 
10759 
diff
changeset
 | 
231  | 
may repeatedly subtract assumptions until none are left!  | 
| 60530 | 232  | 
\<close>  | 
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
233  | 
|
| 77062 | 234  | 
lemma completeness_0:  | 
235  | 
  assumes "{} \<Turnstile> p"
 | 
|
236  | 
  shows "{} \<turnstile> p"
 | 
|
237  | 
proof -  | 
|
238  | 
  { fix t t0
 | 
|
239  | 
have "hyps p t - hyps p t0 \<turnstile> p"  | 
|
240  | 
using hyps_finite hyps_subset  | 
|
241  | 
proof (induction arbitrary: t rule: finite_subset_induct)  | 
|
242  | 
case empty  | 
|
243  | 
then show ?case  | 
|
244  | 
by (simp add: assms sat_thms_p)  | 
|
245  | 
next  | 
|
246  | 
case (insert q H)  | 
|
247  | 
then consider v where "q = #v" | v where "q = #v \<rightharpoonup> false"  | 
|
248  | 
by blast  | 
|
249  | 
then show ?case  | 
|
250  | 
proof cases  | 
|
251  | 
case 1  | 
|
252  | 
then show ?thesis  | 
|
253  | 
by (metis (no_types, lifting) insert.IH thms_excluded_middle_rule insert_Diff_same  | 
|
254  | 
insert_Diff_subset2 weaken_left Diff_weaken_left hyps_Diff)  | 
|
255  | 
next  | 
|
256  | 
case 2  | 
|
257  | 
then show ?thesis  | 
|
258  | 
by (metis (no_types, lifting) insert.IH thms_excluded_middle_rule insert_Diff_same  | 
|
259  | 
insert_Diff_subset2 weaken_left Diff_weaken_left hyps_insert)  | 
|
260  | 
qed  | 
|
261  | 
qed  | 
|
262  | 
}  | 
|
263  | 
then show ?thesis  | 
|
264  | 
by (metis Diff_cancel)  | 
|
265  | 
qed  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
266  | 
|
| 60530 | 267  | 
text\<open>A semantic analogue of the Deduction Theorem\<close>  | 
| 77062 | 268  | 
lemma sat_imp: "insert p H \<Turnstile> q \<Longrightarrow> H \<Turnstile> p\<rightharpoonup>q"  | 
269  | 
by (auto simp: sat_def)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
270  | 
|
| 77062 | 271  | 
theorem completeness: "finite H \<Longrightarrow> H \<Turnstile> p \<Longrightarrow> H \<turnstile> p"  | 
272  | 
proof (induction arbitrary: p rule: finite_induct)  | 
|
273  | 
case empty  | 
|
274  | 
then show ?case  | 
|
275  | 
by (simp add: completeness_0)  | 
|
276  | 
next  | 
|
277  | 
case insert  | 
|
278  | 
then show ?case  | 
|
279  | 
by (meson H MP insertI1 sat_imp weaken_left_insert)  | 
|
280  | 
qed  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
281  | 
|
| 77062 | 282  | 
theorem syntax_iff_semantics: "finite H \<Longrightarrow> (H \<turnstile> p) = (H \<Turnstile> p)"  | 
283  | 
by (blast intro: soundness completeness)  | 
|
| 
13075
 
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
 
paulson 
parents: 
10759 
diff
changeset
 | 
284  | 
|
| 
3120
 
c58423c20740
New directory to contain examples of (co)inductive definitions
 
paulson 
parents:  
diff
changeset
 | 
285  | 
end  |