src/ZF/ex/PropLog.thy
author wenzelm
Sat, 20 Oct 2001 20:20:41 +0200
changeset 11852 a528a716a312
parent 11354 9b80fe19407f
permissions -rw-r--r--
added TextIO.stdErr;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/ex/PropLog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Tobias Nipkow & Lawrence C Paulson
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
     6
Datatype definition of propositional logic formulae and inductive definition
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
     7
of the propositional tautologies.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
11354
9b80fe19407f examples files start from Main instead of various ZF theories
paulson
parents: 11316
diff changeset
    10
PropLog = Main +
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    11
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    12
(** The datatype of propositions; note mixfix syntax **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    14
  prop     :: i
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    16
datatype
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    17
  "prop" = Fls
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    18
         | Var ("n \\<in> nat")                       ("#_" [100] 100)
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    19
         | "=>" ("p \\<in> prop", "q \\<in> prop")          (infixr 90)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    20
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    21
(** The proof system **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    22
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    23
  thms     :: i => i
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    24
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    25
syntax
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    26
  "|-"     :: [i,i] => o                        (infixl 50)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
translations
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    29
  "H |- p" == "p \\<in> thms(H)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    31
inductive
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    32
  domains "thms(H)" <= "prop"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    33
  intrs
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    34
    H  "[| p \\<in> H;  p \\<in> prop |] ==> H |- p"
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    35
    K  "[| p \\<in> prop;  q \\<in> prop |] ==> H |- p=>q=>p"
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    36
    S  "[| p \\<in> prop;  q \\<in> prop;  r \\<in> prop |] ==> H |- (p=>q=>r) => (p=>q) => p=>r"
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    37
    DN "p \\<in> prop ==> H |- ((p=>Fls) => Fls) => p"
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    38
    MP "[| H |- p=>q;  H |- p;  p \\<in> prop;  q \\<in> prop |] ==> H |- q"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    39
  type_intrs "prop.intrs"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    40
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    41
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    42
(** The semantics **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    43
consts
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    44
  "|="        :: [i,i] => o                        (infixl 50)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    45
  hyps        :: [i,i] => i
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    46
  is_true_fun :: [i,i] => i
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    47
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    48
constdefs (*this definitionis necessary since predicates can't be recursive*)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    49
  is_true     :: [i,i] => o
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    50
    "is_true(p,t) == is_true_fun(p,t)=1"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 501
diff changeset
    51
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 515
diff changeset
    52
defs
501
9c724f7085f9 ZF/ex/PropLog/sat_XXX: renamed logcon_XXX, since the relation is logical
lcp
parents: 0
diff changeset
    53
  (*Logical consequence: for every valuation, if all elements of H are true
9c724f7085f9 ZF/ex/PropLog/sat_XXX: renamed logcon_XXX, since the relation is logical
lcp
parents: 0
diff changeset
    54
     then so is p*)
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    55
  logcon_def  "H |= p == \\<forall>t. (\\<forall>q \\<in> H. is_true(q,t)) --> is_true(p,t)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    57
primrec (** A finite set of hypotheses from t and the Vars in p **)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    58
  "hyps(Fls, t)    = 0"
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    59
  "hyps(Var(v), t) = (if v \\<in> t then {#v} else {#v=>Fls})"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    60
  "hyps(p=>q, t)   = hyps(p,t) Un hyps(q,t)"
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    61
 
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    62
primrec (** Semantics of propositional logic **)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    63
  "is_true_fun(Fls, t)    = 0"
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6068
diff changeset
    64
  "is_true_fun(Var(v), t) = (if v \\<in> t then 1 else 0)"
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 6046
diff changeset
    65
  "is_true_fun(p=>q, t)   = (if is_true_fun(p,t)=1 then is_true_fun(q,t)
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 6046
diff changeset
    66
			     else 1)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
end