(* Title: HOL/HOL.thy ID: $Id: HOL_BASIC.thy,v 1.2 2003/02/28 13:30:47 smaus Exp $ Author: Tobias Nipkow Copyright 1993 University of Cambridge Higher-Order Logic. *) theory HOL_BASIC = CPure: (*Jan Smaus removed some calls to files here that contain lemmas that are subject of the exercises *) (** Core syntax **) global classes "term" < logic defaultsort "term" typedecl bool arities bool :: "term" fun :: ("term", "term") "term" consts (* Constants *) Trueprop :: "bool => prop" ("(_)" 5) Not :: "bool => bool" ("~ _" [40] 40) True :: bool False :: bool If :: "[bool, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) arbitrary :: 'a (* Binders *) Eps :: "('a => bool) => 'a" All :: "('a => bool) => bool" (binder "ALL " 10) Ex :: "('a => bool) => bool" (binder "EX " 10) Ex1 :: "('a => bool) => bool" (binder "EX! " 10) Let :: "['a, 'a => 'b] => 'b" (* Infixes *) "=" :: "['a, 'a] => bool" (infixl 50) & :: "[bool, bool] => bool" (infixr 35) "|" :: "[bool, bool] => bool" (infixr 30) --> :: "[bool, bool] => bool" (infixr 25) local (** Additional concrete syntax **) nonterminals letbinds letbind case_syn cases_syn syntax ~= :: "['a, 'a] => bool" (infixl 50) "_Eps" :: "[pttrn, bool] => 'a" ("(3SOME _./ _)" [0, 10] 10) (* Let expressions *) "_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10) "" :: "letbind => letbinds" ("_") "_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") "_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) (* Case expressions *) "_case_syntax":: "['a, cases_syn] => 'b" ("(case _ of/ _)" 10) "_case1" :: "['a, 'b] => case_syn" ("(2_ =>/ _)" 10) "" :: "case_syn => cases_syn" ("_") "_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ | _") translations "x ~= y" == "~ (x = y)" "SOME x. P" == "Eps (%x. P)" "_Let (_binds b bs) e" == "_Let b (_Let bs e)" "let x = a in e" == "Let a (%x. e)" syntax ("" output) "op =" :: "['a, 'a] => bool" ("(_ =/ _)" [51, 51] 50) "op ~=" :: "['a, 'a] => bool" ("(_ ~=/ _)" [51, 51] 50) syntax (symbols) Not :: "bool => bool" ("\\ _" [40] 40) "op &" :: "[bool, bool] => bool" (infixr "\\" 35) "op |" :: "[bool, bool] => bool" (infixr "\\" 30) "op -->" :: "[bool, bool] => bool" (infixr "\\\\" 25) "op ~=" :: "['a, 'a] => bool" (infixl "\\" 50) "ALL " :: "[idts, bool] => bool" ("(3\\_./ _)" [0, 10] 10) "EX " :: "[idts, bool] => bool" ("(3\\_./ _)" [0, 10] 10) "EX! " :: "[idts, bool] => bool" ("(3\\!_./ _)" [0, 10] 10) "_case1" :: "['a, 'b] => case_syn" ("(2_ \\/ _)" 10) (*"_case2" :: "[case_syn, cases_syn] => cases_syn" ("_/ \\ _")*) syntax (input) "_Eps" :: "[pttrn, bool] => 'a" ("(3\\_./ _)" [0, 10] 10) syntax (symbols output) "op ~=" :: "['a, 'a] => bool" ("(_ \\/ _)" [51, 51] 50) syntax (xsymbols) "op -->" :: "[bool, bool] => bool" (infixr "\\" 25) syntax (HOL) "_Eps" :: "[pttrn, bool] => 'a" ("(3@ _./ _)" [0, 10] 10) "ALL " :: "[idts, bool] => bool" ("(3! _./ _)" [0, 10] 10) "EX " :: "[idts, bool] => bool" ("(3? _./ _)" [0, 10] 10) "EX! " :: "[idts, bool] => bool" ("(3?! _./ _)" [0, 10] 10) (** Rules and definitions **) axioms eq_reflection: "(x=y) ==> (x==y)" (* Basic Rules *) refl: "t = (t::'a)" subst: "[| s = t; P(s) |] ==> P(t::'a)" (*Extensionality is built into the meta-logic, and this rule expresses a related property. It is an eta-expanded version of the traditional rule, and similar to the ABS rule of HOL.*) ext: "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)" someI: "P (x::'a) ==> P (SOME x. P x)" impI: "(P ==> Q) ==> P-->Q" mp: "[| P-->Q; P |] ==> Q" defs True_def: "True == ((%x::bool. x) = (%x. x))" All_def: "All(P) == (P = (%x. True))" Ex_def: "Ex(P) == P (SOME x. P x)" False_def: "False == (!P. P)" not_def: "~ P == P-->False" and_def: "P & Q == !R. (P-->Q-->R) --> R" or_def: "P | Q == !R. (P-->R) --> (Q-->R) --> R" Ex1_def: "Ex1(P) == ? x. P(x) & (! y. P(y) --> y=x)" axioms (* Axioms *) iff: "(P-->Q) --> (Q-->P) --> (P=Q)" True_or_False: "(P=True) | (P=False)" defs (*misc definitions*) Let_def: "Let s f == f(s)" if_def: "If P x y == SOME z::'a. (P=True --> z=x) & (P=False --> z=y)" (*arbitrary is completely unspecified, but is made to appear as a definition syntactically*) arbitrary_def: "False ==> arbitrary == (SOME x. False)" end