author | convert-repo |
Thu, 23 Jul 2009 14:03:20 +0000 | |
changeset 255 | 435bf30c29a5 |
parent 102 | 18d44ab74672 |
permissions | -rw-r--r-- |
56
385d51d74f71
Used Datatype functor to define propositional logic terms.
nipkow
parents:
48
diff
changeset
|
1 |
(* Title: HOL/ex/pl.thy |
0 | 2 |
ID: $Id$ |
3 |
Author: Tobias Nipkow |
|
56
385d51d74f71
Used Datatype functor to define propositional logic terms.
nipkow
parents:
48
diff
changeset
|
4 |
Copyright 1994 TU Muenchen |
0 | 5 |
|
6 |
Inductive definition of propositional logic. |
|
7 |
*) |
|
8 |
||
93 | 9 |
PL = Finite + |
10 |
datatype |
|
11 |
'a pl = false | var ('a) ("#_" [1000]) | "->" ('a pl,'a pl) (infixr 90) |
|
0 | 12 |
consts |
13 |
axK,axS,axDN:: "'a pl set" |
|
14 |
ruleMP,thms :: "'a pl set => 'a pl set" |
|
15 |
"|-" :: "['a pl set, 'a pl] => bool" (infixl 50) |
|
16 |
"|=" :: "['a pl set, 'a pl] => bool" (infixl 50) |
|
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
17 |
eval2 :: "['a pl, 'a set] => bool" |
0 | 18 |
eval :: "['a set, 'a pl] => bool" ("_[_]" [100,0] 100) |
19 |
hyps :: "['a pl, 'a set] => 'a pl set" |
|
20 |
rules |
|
21 |
||
22 |
(** Proof theory for propositional logic **) |
|
23 |
||
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
24 |
axK_def "axK == {x . ? p q. x = p->q->p}" |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
25 |
axS_def "axS == {x . ? p q r. x = (p->q->r) -> (p->q) -> p->r}" |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
26 |
axDN_def "axDN == {x . ? p. x = ((p->false) -> false) -> p}" |
0 | 27 |
|
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
28 |
(*the use of subsets simplifies the proof of monotonicity*) |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
29 |
ruleMP_def "ruleMP(X) == {q. ? p:X. p->q : X}" |
0 | 30 |
|
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
31 |
thms_def |
0 | 32 |
"thms(H) == lfp(%X. H Un axK Un axS Un axDN Un ruleMP(X))" |
33 |
||
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
34 |
conseq_def "H |- p == p : thms(H)" |
0 | 35 |
|
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
36 |
sat_def "H |= p == (!tt. (!q:H. tt[q]) --> tt[p])" |
0 | 37 |
|
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
38 |
eval_def "tt[p] == eval2(p,tt)" |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
39 |
primrec eval2 pl |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
40 |
eval2_false "eval2(false) = (%x.False)" |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
41 |
eval2_var "eval2(#v) = (%tt.v:tt)" |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
42 |
eval2_imp "eval2(p->q) = (%tt.eval2(p,tt)-->eval2(q,tt))" |
0 | 43 |
|
102
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
44 |
primrec hyps pl |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
45 |
hyps_false "hyps(false) = (%tt.{})" |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
46 |
hyps_var "hyps(#v) = (%tt.{if(v:tt, #v, #v->false)})" |
18d44ab74672
Used the new primitive recursive functions format for thy-files
nipkow
parents:
93
diff
changeset
|
47 |
hyps_imp "hyps(p->q) = (%tt.hyps(p,tt) Un hyps(q,tt))" |
56
385d51d74f71
Used Datatype functor to define propositional logic terms.
nipkow
parents:
48
diff
changeset
|
48 |
|
0 | 49 |
end |