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