author | paulson |
Thu, 26 Nov 1998 17:40:10 +0100 | |
changeset 5977 | 9f0c8869cf71 |
parent 5310 | 3e14d6d66dab |
child 6260 | a8010d459ef7 |
permissions | -rw-r--r-- |
1459 | 1 |
(* Title: FOL/ROOT |
0 | 2 |
ID: $Id$ |
1459 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
6 |
Adds First-Order Logic to a database containing pure Isabelle. |
|
7 |
Should be executed in the subdirectory FOL. |
|
8 |
*) |
|
9 |
||
10 |
val banner = "First-Order Logic with Natural Deduction"; |
|
11 |
||
12 |
writeln banner; |
|
13 |
||
14 |
print_depth 1; |
|
15 |
||
4222 | 16 |
use "$ISABELLE_HOME/src/Provers/simplifier.ML"; |
17 |
use "$ISABELLE_HOME/src/Provers/splitter.ML"; |
|
18 |
use "$ISABELLE_HOME/src/Provers/ind.ML"; |
|
19 |
use "$ISABELLE_HOME/src/Provers/hypsubst.ML"; |
|
20 |
use "$ISABELLE_HOME/src/Provers/classical.ML"; |
|
21 |
use "$ISABELLE_HOME/src/Provers/blast.ML"; |
|
5219 | 22 |
use "$ISABELLE_HOME/src/Provers/clasimp.ML"; |
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4223
diff
changeset
|
23 |
use "$ISABELLE_HOME/src/Provers/quantifier1.ML"; |
0 | 24 |
|
2469 | 25 |
use_thy "IFOL"; |
4349
50403e5a44c0
Instantiated the one-point-rule quantifier simpprocs for FOL
paulson
parents:
4223
diff
changeset
|
26 |
use "fologic.ML"; |
0 | 27 |
|
2866 | 28 |
(** Applying HypsubstFun to generate hyp_subst_tac **) |
0 | 29 |
structure Hypsubst_Data = |
30 |
struct |
|
1004 | 31 |
structure Simplifier = Simplifier |
4466
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
paulson
parents:
4349
diff
changeset
|
32 |
(*These destructors Match!*) |
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
paulson
parents:
4349
diff
changeset
|
33 |
fun dest_eq (Const("op =",T) $ t $ u) = (t, u, domain_type T) |
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
paulson
parents:
4349
diff
changeset
|
34 |
val dest_Trueprop = FOLogic.dest_Trueprop |
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
paulson
parents:
4349
diff
changeset
|
35 |
val dest_imp = FOLogic.dest_imp |
1004 | 36 |
val eq_reflection = eq_reflection |
0 | 37 |
val imp_intr = impI |
38 |
val rev_mp = rev_mp |
|
39 |
val subst = subst |
|
40 |
val sym = sym |
|
4223 | 41 |
val thin_refl = prove_goal IFOL.thy |
42 |
"!!X. [|x=x; PROP W|] ==> PROP W" (K [atac 1]); |
|
0 | 43 |
end; |
44 |
||
45 |
structure Hypsubst = HypsubstFun(Hypsubst_Data); |
|
46 |
open Hypsubst; |
|
47 |
||
2469 | 48 |
|
97 | 49 |
use "intprover.ML"; |
0 | 50 |
|
2469 | 51 |
use_thy "FOL"; |
0 | 52 |
|
4096 | 53 |
use "cladata.ML"; |
54 |
use "simpdata.ML"; |
|
55 |
||
1523 | 56 |
qed_goal "ex1_functional" FOL.thy |
666 | 57 |
"!!a b c. [| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c" |
5310 | 58 |
(fn _ => [ (Blast_tac 1) ]); |
0 | 59 |
|
60 |
||
61 |
print_depth 8; |
|
62 |
||
1459 | 63 |
val FOL_build_completed = (); (*indicate successful build*) |