improved arrangement of files;
authorwenzelm
Thu, 29 Jun 2000 22:48:08 +0200
changeset 9205 f171fa6a0989
parent 9204 e865dda0313e
child 9206 eaaee6bd74ba
improved arrangement of files; tuned;
src/FOL/IsaMakefile
src/FOL/ex/If.ML
src/FOL/ex/ROOT.ML
src/FOL/ex/foundn.ML
src/FOL/ex/int.ML
src/FOL/ex/int.thy
--- a/src/FOL/IsaMakefile	Thu Jun 29 22:39:57 2000 +0200
+++ b/src/FOL/IsaMakefile	Thu Jun 29 22:48:08 2000 +0200
@@ -43,7 +43,7 @@
 $(LOG)/FOL-ex.gz: $(OUT)/FOL ex/If.ML ex/If.thy ex/IffOracle.ML \
   ex/IffOracle.thy ex/List.ML ex/List.thy ex/Nat.ML ex/Nat.thy \
   ex/Nat2.ML ex/Nat2.thy ex/Prolog.ML ex/Prolog.thy ex/ROOT.ML ex/cla.ML \
-  ex/foundn.ML ex/int.ML ex/intro.ML ex/prop.ML ex/quant.ML
+  ex/foundn.ML ex/int.ML ex/int.thy ex/intro.ML ex/prop.ML ex/quant.ML
 	@$(ISATOOL) usedir $(OUT)/FOL ex
 
 
--- a/src/FOL/ex/If.ML	Thu Jun 29 22:39:57 2000 +0200
+++ b/src/FOL/ex/If.ML	Thu Jun 29 22:48:08 2000 +0200
@@ -1,12 +1,11 @@
-(*  Title:      FOL/ex/if
+(*  Title:      FOL/ex/If.ML
     ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1991  University of Cambridge
 
-First-Order Logic: the 'if' example
+First-Order Logic: the 'if' example.
 *)
 
-open If;
 open Cla;    (*in case structure IntPr is open!*)
 
 val prems = Goalw [if_def]
@@ -42,7 +41,7 @@
 choplev 0;
 by (rewtac if_def);
 by (blast_tac FOL_cs 1);
-result();
+qed "";
 
 
 (*An invalid formula.  High-level rules permit a simpler diagnosis*)
@@ -58,7 +57,3 @@
 (*Check that subgoals remain: proof failed.*)
 getgoal 1; 
 by (REPEAT (step_tac FOL_cs 1));
-
-
-
-writeln"Reached end of file.";
--- a/src/FOL/ex/ROOT.ML	Thu Jun 29 22:39:57 2000 +0200
+++ b/src/FOL/ex/ROOT.ML	Thu Jun 29 22:48:08 2000 +0200
@@ -12,7 +12,7 @@
 time_use_thy "Prolog";
 
 writeln"\n** Intuitionistic examples **\n";
-time_use     "int.ML";
+time_use_thy "int";
 
 val thy = IFOL.thy  and  tac = IntPr.fast_tac 1;
 time_use     "prop.ML";
--- a/src/FOL/ex/foundn.ML	Thu Jun 29 22:39:57 2000 +0200
+++ b/src/FOL/ex/foundn.ML	Thu Jun 29 22:48:08 2000 +0200
@@ -6,8 +6,6 @@
 Intuitionistic FOL: Examples from The Foundation of a Generic Theorem Prover
 *)
 
-writeln"File FOL/ex/foundn.ML";
-
 goal IFOL.thy "A&B  --> (C-->A&C)";
 by (rtac impI 1);
 by (rtac impI 1);
@@ -129,6 +127,3 @@
 by (rtac allE 1 THEN assume_tac 1);
 by (assume_tac 1);
 result();  
-
-
-writeln"Reached end of file.";
--- a/src/FOL/ex/int.ML	Thu Jun 29 22:39:57 2000 +0200
+++ b/src/FOL/ex/int.ML	Thu Jun 29 22:48:08 2000 +0200
@@ -1,9 +1,9 @@
-(*  Title:      FOL/ex/int
+(*  Title:      FOL/ex/int.ML
     ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1991  University of Cambridge
 
-Intuitionistic First-Order Logic
+Intuitionistic First-Order Logic.
 
 Single-step commands:
 by (IntPr.step_tac 1);
@@ -15,9 +15,6 @@
 by (IntPr.fast_tac 1);
 *)
 
-context IFOL.thy;
-
-writeln"File FOL/ex/int.";
 
 (*Metatheorem (for PROPOSITIONAL formulae...):
   P is classically provable iff ~~P is intuitionistically provable.
@@ -33,54 +30,54 @@
 
 Goal "~~(P&Q) <-> ~~P & ~~Q";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "~~ ((~P --> Q) --> (~P --> ~Q) --> P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 (* ~~ does NOT distribute over | *)
 
 Goal "~~(P-->Q)  <-> (~~P --> ~~Q)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "~~~P <-> ~P";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "~~((P --> Q | R)  -->  (P-->Q) | (P-->R))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "(P<->Q) <-> (Q<->P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "((P --> (Q | (Q-->R))) --> R) --> R";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "(((G-->A) --> J) --> D --> E) --> (((H-->B)-->I)-->C-->J) \
 \     --> (A-->H) --> F --> G --> (((C-->B)-->I)-->D)-->(A-->C) \
 \     --> (((F-->A)-->B) --> I) --> E";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 
 writeln"Lemmas for the propositional double-negation translation";
 
 Goal "P --> ~~P";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "~~(~~P --> P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "~~P & ~~(P --> Q) --> ~~Q";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 
 writeln"The following are classically but not constructively valid.";
@@ -103,7 +100,7 @@
 \              ((Q<->R) --> P&Q&R) &    \
 \              ((R<->P) --> P&Q&R) --> P&Q&R";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 
 (*de Bruijn formula with five predicates*)
@@ -113,7 +110,7 @@
 \              ((S<->T) --> P&Q&R&S&T) &    \
 \              ((T<->P) --> P&Q&R&S&T) --> P&Q&R&S&T";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 
 (*** Problems from of Sahlin, Franzen and Haridi, 
@@ -131,7 +128,7 @@
 (*Problem 3.1*)
 Goal "~ (EX x. ALL y. mem(y,x) <-> ~ mem(x,x))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 (*Problem 4.1: hopeless!*)
 Goal "(ALL x. p(x) --> p(h(x)) | p(g(x))) & (EX x. p(x)) & (ALL x. ~p(h(x))) \
@@ -143,60 +140,60 @@
 writeln"Problem ~~1";
 Goal "~~((P-->Q)  <->  (~Q --> ~P))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 
 writeln"Problem ~~2";
 Goal "~~(~~P  <->  P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 (*1 secs*)
 
 
 writeln"Problem 3";
 Goal "~(P-->Q) --> (Q-->P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~4";
 Goal "~~((~P-->Q)  <->  (~Q --> P))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 (*9 secs*)
 
 writeln"Problem ~~5";
 Goal "~~((P|Q-->P|R) --> P|(Q-->R))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 (*10 secs*)
 
 
 writeln"Problem ~~6";
 Goal "~~(P | ~P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~7";
 Goal "~~(P | ~~~P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~8.  Peirce's law";
 Goal "~~(((P-->Q) --> P)  -->  P)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 9";
 Goal "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 (*9 secs*)
 
 
 writeln"Problem 10";
 Goal "(Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"11.  Proved in each direction (incorrectly, says Pelletier!!) ";
 Goal "P<->P";
@@ -205,42 +202,42 @@
 writeln"Problem ~~12.  Dijkstra's law  ";
 Goal "~~(((P <-> Q) <-> R)  <->  (P <-> (Q <-> R)))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 Goal "((P <-> Q) <-> R)  -->  ~~(P <-> (Q <-> R))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 13.  Distributive law";
 Goal "P | (Q & R)  <-> (P | Q) & (P | R)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~14";
 Goal "~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~15";
 Goal "~~((P --> Q) <-> (~P | Q))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~16";
 Goal "~~((P-->Q) | (Q-->P))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~17";
 Goal
   "~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))";
 by (IntPr.fast_tac 1);    
-result();
+qed "";
 
 (*Dijkstra's "Golden Rule"*)
 Goal "(P&Q) <-> P <-> Q <-> (P|Q)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 
 writeln"****Examples with quantifiers****";
@@ -250,23 +247,23 @@
 
 Goal "(EX x. P(x)-->Q)  -->  (ALL x. P(x)) --> Q";
 by (IntPr.fast_tac 1); 
-result();  
+qed "";  
 
 Goal "((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)";
 by (IntPr.fast_tac 1); 
-result();  
+qed "";  
 
 Goal "((ALL x. ~P(x))-->Q)  -->  ~ (ALL x. ~ (P(x)|Q))";
 by (IntPr.fast_tac 1); 
-result();  
+qed "";  
 
 Goal "(ALL x. P(x)) | Q  -->  (ALL x. P(x) | Q)";
 by (IntPr.fast_tac 1); 
-result();  
+qed "";  
 
 Goal "(EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))";
 by (IntPr.fast_tac 1);
-result();  
+qed "";  
 
 
 
@@ -313,7 +310,7 @@
 Goal "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))     \
 \   --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))";
 by (IntPr.fast_tac 1); 
-result();
+qed "";
 
 writeln"Problem 21";
 Goal "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))";
@@ -322,12 +319,12 @@
 writeln"Problem 22";
 Goal "(ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))";
 by (IntPr.fast_tac 1); 
-result();
+qed "";
 
 writeln"Problem ~~23";
 Goal "~~ ((ALL x. P | Q(x))  <->  (P | (ALL x. Q(x))))";
 by (IntPr.fast_tac 1); 
-result();
+qed "";
 
 writeln"Problem 24";
 Goal "~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &  \
@@ -338,7 +335,7 @@
 by (etac impE 1);
 by (IntPr.fast_tac 1); 
 by (IntPr.fast_tac 1); 
-result();
+qed "";
 
 writeln"Problem 25";
 Goal "(EX x. P(x)) &  \
@@ -347,7 +344,7 @@
 \       ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))  \
 \   --> (EX x. Q(x)&P(x))";
 by (IntPr.fast_tac 1); 
-result();
+qed "";
 
 writeln"Problem ~~26";
 Goal "(~~(EX x. p(x)) <-> ~~(EX x. q(x))) &    \
@@ -362,7 +359,7 @@
 \             ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x)))  \
 \         --> (ALL x. M(x) --> ~L(x))";
 by (IntPr.fast_tac 1);   (*21 secs*)
-result();
+qed "";
 
 writeln"Problem ~~28.  AMENDED";
 Goal "(ALL x. P(x) --> (ALL x. Q(x))) &   \
@@ -370,21 +367,21 @@
 \       (~~(EX x. S(x)) --> (ALL x. L(x) --> M(x)))  \
 \   --> (ALL x. P(x) & L(x) --> M(x))";
 by (IntPr.fast_tac 1);  (*48 secs*)
-result();
+qed "";
 
 writeln"Problem 29.  Essentially the same as Principia Mathematica *11.71";
 Goal "(EX x. P(x)) & (EX y. Q(y))  \
 \   --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y))   <->     \
 \        (ALL x y. P(x) & Q(y) --> R(x) & S(y)))";
 by (IntPr.fast_tac 1); 
-result();
+qed "";
 
 writeln"Problem ~~30";
 Goal "(ALL x. (P(x) | Q(x)) --> ~ R(x)) & \
 \       (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))  \
 \   --> (ALL x. ~~S(x))";
 by (IntPr.fast_tac 1);  
-result();
+qed "";
 
 writeln"Problem 31";
 Goal "~(EX x. P(x) & (Q(x) | R(x))) & \
@@ -392,7 +389,7 @@
 \       (ALL x. ~ R(x) --> M(x))  \
 \   --> (EX x. L(x) & M(x))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 32";
 Goal "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) & \
@@ -400,13 +397,13 @@
 \       (ALL x. M(x) --> R(x))  \
 \   --> (ALL x. P(x) & M(x) --> L(x))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem ~~33";
 Goal "(ALL x. ~~(P(a) & (P(x)-->P(b))-->P(c)))  <->    \
 \     (ALL x. ~~((~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c))))";
 by (IntPr.best_tac 1);   (*1.67s*)
-result();
+qed "";
 
 
 writeln"Problem 36";
@@ -416,7 +413,7 @@
 \     (ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z)))   \
 \ --> (ALL x. EX y. H(x,y))";
 by (IntPr.fast_tac 1);  (*5 secs*)
-result();
+qed "";
 
 writeln"Problem 37";
 Goal
@@ -430,13 +427,13 @@
 writeln"Problem 39";
 Goal "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 40.  AMENDED";
 Goal "(EX y. ALL x. F(x,y) <-> F(x,x)) -->  \
 \             ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 44";
 Goal "(ALL x. f(x) -->                                 \
@@ -444,40 +441,38 @@
 \             (EX x. j(x) & (ALL y. g(y) --> h(x,y)))                   \
 \             --> (EX x. j(x) & ~f(x))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 48";
 Goal "(a=b | c=d) & (a=c | b=d) --> a=d | b=c";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 51";
 Goal "(EX z w. ALL x y. P(x,y) <->  (x=z & y=w)) -->  \
 \    (EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 52";
 (*Almost the same as 51. *)
 Goal "(EX z w. ALL x y. P(x,y) <->  (x=z & y=w)) -->  \
 \    (EX w. ALL y. EX z. (ALL x. P(x,y) <-> x=z) <-> y=w)";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 56";
 Goal "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 57";
 Goal "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) & \
 \    (ALL x y z. P(x,y) & P(y,z) --> P(x,z))    -->   P(f(a,b), f(a,c))";
 by (IntPr.fast_tac 1);
-result();
+qed "";
 
 writeln"Problem 60";
 Goal "ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))";
 by (IntPr.fast_tac 1);
-result();
-
-writeln"Reached end of file.";
+qed "";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/FOL/ex/int.thy	Thu Jun 29 22:48:08 2000 +0200
@@ -0,0 +1,11 @@
+(*  Title:      FOL/ex/int.thy
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1991  University of Cambridge
+
+Intuitionistic First-Order Logic.
+*)
+
+theory int = IFOL:
+
+end