--- a/src/FOL/FOL.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/FOL.ML Fri Jan 03 15:01:55 1997 +0100
@@ -67,3 +67,9 @@
[ (rtac conjE 1),
(REPEAT (DEPTH_SOLVE_1
(etac impCE 1 ORELSE mp_tac 1 ORELSE ares_tac prems 1))) ]);
+
+
+(*Thus, assignments to the references claset and simpset are recorded
+ with theory "FOL". These files cannot be loaded directly in ROOT.ML.*)
+use "cladata.ML";
+use "simpdata.ML";
--- a/src/FOL/Makefile Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/Makefile Fri Jan 03 15:01:55 1997 +0100
@@ -22,6 +22,7 @@
BIN = $(ISABELLEBIN)
COMP = $(ISABELLECOMP)
FILES = ROOT.ML IFOL.thy IFOL.ML FOL.thy FOL.ML intprover.ML simpdata.ML \
+ thy_data.ML cladata.ML \
../Provers/hypsubst.ML ../Provers/classical.ML \
../Provers/simplifier.ML ../Provers/splitter.ML ../Provers/ind.ML
--- a/src/FOL/ROOT.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ROOT.ML Fri Jan 03 15:01:55 1997 +0100
@@ -14,14 +14,13 @@
init_thy_reader();
print_depth 1;
-use_thy "FOL";
use "../Provers/splitter.ML";
use "../Provers/ind.ML";
use "../Provers/hypsubst.ML";
use "../Provers/classical.ML";
-(*** Applying HypsubstFun to generate hyp_subst_tac ***)
+use_thy "IFOL";
structure Hypsubst_Data =
struct
@@ -38,36 +37,15 @@
structure Hypsubst = HypsubstFun(Hypsubst_Data);
open Hypsubst;
+
use "intprover.ML";
-(*** Applying ClassicalFun to create a classical prover ***)
-structure Classical_Data =
- struct
- val sizef = size_of_thm
- val mp = mp
- val not_elim = notE
- val classical = classical
- val hyp_subst_tacs=[hyp_subst_tac]
- end;
-
-structure Cla = ClassicalFun(Classical_Data);
-open Cla;
-
-(*Propositional rules
- -- iffCE might seem better, but in the examples in ex/cla
- run about 7% slower than with iffE*)
-val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI]
- addSEs [conjE,disjE,impCE,FalseE,iffE];
-
-(*Quantifier rules*)
-val FOL_cs = prop_cs addSIs [allI] addIs [exI,ex1I]
- addSEs [exE,ex1E] addEs [allE];
+use_thy "FOL";
qed_goal "ex1_functional" FOL.thy
"!!a b c. [| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c"
(fn _ => [ (deepen_tac FOL_cs 0 1) ]);
-use "simpdata.ML";
init_pps ();
print_depth 8;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/FOL/cladata.ML Fri Jan 03 15:01:55 1997 +0100
@@ -0,0 +1,51 @@
+(* Title: FOL/cladata.ML
+ ID: $Id$
+ Author: Tobias Nipkow
+ Copyright 1996 University of Cambridge
+
+Setting up the classical reasoner
+*)
+
+
+(** Applying HypsubstFun to generate hyp_subst_tac **)
+section "Classical Reasoner";
+
+(*** Applying ClassicalFun to create a classical prover ***)
+structure Classical_Data =
+ struct
+ val sizef = size_of_thm
+ val mp = mp
+ val not_elim = notE
+ val classical = classical
+ val hyp_subst_tacs=[hyp_subst_tac]
+ end;
+
+structure Cla = ClassicalFun(Classical_Data);
+open Cla;
+
+(*Propositional rules
+ -- iffCE might seem better, but in the examples in ex/cla
+ run about 7% slower than with iffE*)
+val prop_cs = empty_cs addSIs [refl,TrueI,conjI,disjCI,impI,notI,iffI]
+ addSEs [conjE,disjE,impCE,FalseE,iffE];
+
+(*Quantifier rules*)
+val FOL_cs = prop_cs addSIs [allI] addIs [exI,ex1I]
+ addSEs [exE,ex1E] addEs [allE];
+
+
+exception CS_DATA of claset;
+
+let fun merge [] = CS_DATA empty_cs
+ | merge cs = let val cs = map (fn CS_DATA x => x) cs;
+ in CS_DATA (foldl merge_cs (hd cs, tl cs)) end;
+
+ fun put (CS_DATA cs) = claset := cs;
+
+ fun get () = CS_DATA (!claset);
+in add_thydata "FOL"
+ ("claset", ThyMethods {merge = merge, put = put, get = get})
+end;
+
+claset := FOL_cs;
+
--- a/src/FOL/ex/If.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ex/If.ML Fri Jan 03 15:01:55 1997 +0100
@@ -30,13 +30,14 @@
by (rtac ifI 1);
choplev 0;
-val if_cs = FOL_cs addSIs [ifI] addSEs[ifE];
-by (fast_tac if_cs 1);
+AddSIs [ifI];
+AddSEs [ifE];
+by (Fast_tac 1);
qed "if_commute";
goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))";
-by (fast_tac if_cs 1);
+by (Fast_tac 1);
qed "nested_ifs";
choplev 0;
@@ -47,10 +48,10 @@
(*An invalid formula. High-level rules permit a simpler diagnosis*)
goal If.thy "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))";
-by (fast_tac if_cs 1) handle ERROR => writeln"Failed, as expected";
+by (Fast_tac 1) handle ERROR => writeln"Failed, as expected";
(*Check that subgoals remain: proof failed.*)
getgoal 1;
-by (REPEAT (step_tac if_cs 1));
+by (REPEAT (Step_tac 1));
choplev 0;
by (rewtac if_def);
--- a/src/FOL/ex/List.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ex/List.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,55 +13,53 @@
by (REPEAT (resolve_tac (prems@[allI,impI]) 1));
qed "list_exh";
-val list_rew_thms = [list_distinct1,list_distinct2,app_nil,app_cons,
- hd_eq,tl_eq,forall_nil,forall_cons,list_free,
- len_nil,len_cons,at_0,at_succ];
-
-val list_ss = nat_ss addsimps list_rew_thms;
+Addsimps [list_distinct1, list_distinct2, app_nil, app_cons,
+ hd_eq, tl_eq, forall_nil, forall_cons, list_free,
+ len_nil, len_cons, at_0, at_succ];
goal List.thy "~l=[] --> hd(l).tl(l) = l";
-by(IND_TAC list_exh (simp_tac list_ss) "l" 1);
+by (IND_TAC list_exh Simp_tac "l" 1);
result();
goal List.thy "(l1++l2)++l3 = l1++(l2++l3)";
-by(IND_TAC list_ind (simp_tac list_ss) "l1" 1);
+by (IND_TAC list_ind Simp_tac "l1" 1);
qed "append_assoc";
goal List.thy "l++[] = l";
-by(IND_TAC list_ind (simp_tac list_ss) "l" 1);
+by (IND_TAC list_ind Simp_tac "l" 1);
qed "app_nil_right";
goal List.thy "l1++l2=[] <-> l1=[] & l2=[]";
-by(IND_TAC list_exh (simp_tac list_ss) "l1" 1);
+by (IND_TAC list_exh Simp_tac "l1" 1);
qed "app_eq_nil_iff";
goal List.thy "forall(l++l',P) <-> forall(l,P) & forall(l',P)";
-by(IND_TAC list_ind (simp_tac list_ss) "l" 1);
+by (IND_TAC list_ind Simp_tac "l" 1);
qed "forall_app";
goal List.thy "forall(l,%x.P(x)&Q(x)) <-> forall(l,P) & forall(l,Q)";
-by(IND_TAC list_ind (simp_tac list_ss) "l" 1);
-by(fast_tac FOL_cs 1);
+by (IND_TAC list_ind Simp_tac "l" 1);
+by (Fast_tac 1);
qed "forall_conj";
goal List.thy "~l=[] --> forall(l,P) <-> P(hd(l)) & forall(tl(l),P)";
-by(IND_TAC list_ind (simp_tac list_ss) "l" 1);
+by (IND_TAC list_ind Simp_tac "l" 1);
qed "forall_ne";
(*** Lists with natural numbers ***)
goal List.thy "len(l1++l2) = len(l1)+len(l2)";
-by (IND_TAC list_ind (simp_tac list_ss) "l1" 1);
+by (IND_TAC list_ind Simp_tac "l1" 1);
qed "len_app";
goal List.thy "i<len(l1) --> at(l1++l2,i) = at(l1,i)";
-by (IND_TAC list_ind (simp_tac list_ss) "l1" 1);
+by (IND_TAC list_ind Simp_tac "l1" 1);
by (REPEAT (rtac allI 1));
by (rtac impI 1);
-by (ALL_IND_TAC nat_exh (asm_simp_tac list_ss) 1);
+by (ALL_IND_TAC nat_exh Asm_simp_tac 1);
qed "at_app1";
goal List.thy "at(l1++(x.l2), len(l1)) = x";
-by (IND_TAC list_ind (simp_tac list_ss) "l1" 1);
+by (IND_TAC list_ind Simp_tac "l1" 1);
qed "at_app_hd2";
--- a/src/FOL/ex/Nat.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ex/Nat.ML Fri Jan 03 15:01:55 1997 +0100
@@ -43,27 +43,27 @@
by (rtac rec_Suc 1);
qed "add_Suc";
-val add_ss = FOL_ss addsimps [add_0, add_Suc];
+Addsimps [add_0, add_Suc];
goal Nat.thy "(k+m)+n = k+(m+n)";
by (res_inst_tac [("n","k")] induct 1);
-by (simp_tac add_ss 1);
-by (asm_simp_tac add_ss 1);
+by (Simp_tac 1);
+by (Asm_simp_tac 1);
qed "add_assoc";
goal Nat.thy "m+0 = m";
by (res_inst_tac [("n","m")] induct 1);
-by (simp_tac add_ss 1);
-by (asm_simp_tac add_ss 1);
+by (Simp_tac 1);
+by (Asm_simp_tac 1);
qed "add_0_right";
goal Nat.thy "m+Suc(n) = Suc(m+n)";
by (res_inst_tac [("n","m")] induct 1);
-by (ALLGOALS (asm_simp_tac add_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "add_Suc_right";
-val [prem] = goal Nat.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)";
+val prems = goal Nat.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)";
by (res_inst_tac [("n","i")] induct 1);
-by (simp_tac add_ss 1);
-by (asm_simp_tac (add_ss addsimps [prem]) 1);
+by (Simp_tac 1);
+by (asm_simp_tac (!simpset addsimps prems) 1);
result();
--- a/src/FOL/ex/Nat2.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ex/Nat2.ML Fri Jan 03 15:01:55 1997 +0100
@@ -9,12 +9,11 @@
open Nat2;
-val nat_rews = [pred_0, pred_succ, plus_0, plus_succ,
- nat_distinct1, nat_distinct2, succ_inject,
- leq_0,leq_succ_succ,leq_succ_0,
- lt_0_succ,lt_succ_succ,lt_0];
+Addsimps [pred_0, pred_succ, plus_0, plus_succ,
+ nat_distinct1, nat_distinct2, succ_inject,
+ leq_0, leq_succ_succ, leq_succ_0,
+ lt_0_succ, lt_succ_succ, lt_0];
-val nat_ss = FOL_ss addsimps nat_rews;
val prems = goal Nat2.thy
"[| P(0); !!x. P(succ(x)) |] ==> All(P)";
@@ -23,143 +22,144 @@
qed "nat_exh";
goal Nat2.thy "~ n=succ(n)";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "~ succ(n)=n";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "~ succ(succ(n))=n";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "~ n=succ(succ(n))";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "m+0 = m";
-by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1);
+by (IND_TAC nat_ind Simp_tac "m" 1);
qed "plus_0_right";
goal Nat2.thy "m+succ(n) = succ(m+n)";
-by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1);
+by (IND_TAC nat_ind Simp_tac "m" 1);
qed "plus_succ_right";
+Addsimps [plus_0_right, plus_succ_right];
+
goal Nat2.thy "~n=0 --> m+pred(n) = pred(m+n)";
-by (IND_TAC nat_ind (simp_tac (nat_ss addsimps [plus_succ_right])) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "~n=0 --> succ(pred(n))=n";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "m+n=0 <-> m=0 & n=0";
-by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1);
+by (IND_TAC nat_ind Simp_tac "m" 1);
result();
goal Nat2.thy "m <= n --> m <= succ(n)";
-by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1);
+by (IND_TAC nat_ind Simp_tac "m" 1);
by (rtac (impI RS allI) 1);
-by (ALL_IND_TAC nat_ind (simp_tac nat_ss) 1);
-by (fast_tac FOL_cs 1);
-val le_imp_le_succ = store_thm("le_imp_le_succ", result() RS mp);
+by (ALL_IND_TAC nat_ind Simp_tac 1);
+by (Fast_tac 1);
+bind_thm("le_imp_le_succ", result() RS mp);
goal Nat2.thy "n<succ(n)";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "~ n<n";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
result();
goal Nat2.thy "m < n --> m < succ(n)";
-by (IND_TAC nat_ind (simp_tac nat_ss) "m" 1);
+by (IND_TAC nat_ind Simp_tac "m" 1);
by (rtac (impI RS allI) 1);
-by (ALL_IND_TAC nat_ind (simp_tac nat_ss) 1);
-by (fast_tac FOL_cs 1);
+by (ALL_IND_TAC nat_ind Simp_tac 1);
+by (Fast_tac 1);
result();
goal Nat2.thy "m <= n --> m <= n+k";
-by (IND_TAC nat_ind
- (simp_tac (nat_ss addsimps [plus_0_right, plus_succ_right, le_imp_le_succ]))
+by (IND_TAC nat_ind (simp_tac (!simpset addsimps [le_imp_le_succ]))
"k" 1);
qed "le_plus";
goal Nat2.thy "succ(m) <= n --> m <= n";
by (res_inst_tac [("x","n")]spec 1);
-by (ALL_IND_TAC nat_exh (simp_tac (nat_ss addsimps [le_imp_le_succ])) 1);
+by (ALL_IND_TAC nat_exh (simp_tac (!simpset addsimps [le_imp_le_succ])) 1);
qed "succ_le";
goal Nat2.thy "~m<n <-> n<=m";
-by (IND_TAC nat_ind (simp_tac nat_ss) "n" 1);
+by (IND_TAC nat_ind Simp_tac "n" 1);
by (rtac (impI RS allI) 1);
-by (ALL_IND_TAC nat_ind (asm_simp_tac nat_ss) 1);
+by (ALL_IND_TAC nat_ind Asm_simp_tac 1);
qed "not_less";
goal Nat2.thy "n<=m --> ~m<n";
-by (simp_tac (nat_ss addsimps [not_less]) 1);
+by (simp_tac (!simpset addsimps [not_less]) 1);
qed "le_imp_not_less";
goal Nat2.thy "m<n --> ~n<=m";
-by (cut_facts_tac [not_less] 1 THEN fast_tac FOL_cs 1);
+by (cut_facts_tac [not_less] 1 THEN Fast_tac 1);
qed "not_le";
goal Nat2.thy "m+k<=n --> m<=n";
by (IND_TAC nat_ind (K all_tac) "k" 1);
-by (simp_tac (nat_ss addsimps [plus_0_right]) 1);
+by (Simp_tac 1);
by (rtac (impI RS allI) 1);
-by (simp_tac (nat_ss addsimps [plus_succ_right]) 1);
+by (Simp_tac 1);
by (REPEAT (resolve_tac [allI,impI] 1));
by (cut_facts_tac [succ_le] 1);
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
qed "plus_le";
val prems = goal Nat2.thy "[| ~m=0; m <= n |] ==> ~n=0";
by (cut_facts_tac prems 1);
by (REPEAT (etac rev_mp 1));
-by (IND_TAC nat_exh (simp_tac nat_ss) "m" 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
+by (IND_TAC nat_exh Simp_tac "m" 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
qed "not0";
goal Nat2.thy "a<=a' & b<=b' --> a+b<=a'+b'";
-by (IND_TAC nat_ind (simp_tac (nat_ss addsimps [plus_0_right,le_plus])) "b" 1);
+by (IND_TAC nat_ind (simp_tac (!simpset addsimps [le_plus])) "b" 1);
by (resolve_tac [impI RS allI] 1);
by (resolve_tac [allI RS allI] 1);
-by (ALL_IND_TAC nat_exh (asm_simp_tac (nat_ss addsimps [plus_succ_right])) 1);
+by (ALL_IND_TAC nat_exh Asm_simp_tac 1);
qed "plus_le_plus";
goal Nat2.thy "i<=j --> j<=k --> i<=k";
by (IND_TAC nat_ind (K all_tac) "i" 1);
-by (simp_tac nat_ss 1);
+by (Simp_tac 1);
by (resolve_tac [impI RS allI] 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
by (resolve_tac [impI] 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
-by (fast_tac FOL_cs 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
+by (Fast_tac 1);
qed "le_trans";
goal Nat2.thy "i < j --> j <=k --> i < k";
by (IND_TAC nat_ind (K all_tac) "j" 1);
-by (simp_tac nat_ss 1);
+by (Simp_tac 1);
by (resolve_tac [impI RS allI] 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
-by (fast_tac FOL_cs 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
+by (Fast_tac 1);
qed "less_le_trans";
goal Nat2.thy "succ(i) <= j <-> i < j";
-by (IND_TAC nat_ind (simp_tac nat_ss) "j" 1);
+by (IND_TAC nat_ind Simp_tac "j" 1);
by (resolve_tac [impI RS allI] 1);
-by (ALL_IND_TAC nat_exh (asm_simp_tac nat_ss) 1);
+by (ALL_IND_TAC nat_exh Asm_simp_tac 1);
qed "succ_le2";
goal Nat2.thy "i<succ(j) <-> i=j | i<j";
-by (IND_TAC nat_ind (simp_tac nat_ss) "j" 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
+by (IND_TAC nat_ind Simp_tac "j" 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
by (resolve_tac [impI RS allI] 1);
-by (ALL_IND_TAC nat_exh (simp_tac nat_ss) 1);
-by (asm_simp_tac nat_ss 1);
+by (ALL_IND_TAC nat_exh Simp_tac 1);
+by (Asm_simp_tac 1);
qed "less_succ";
--- a/src/FOL/ex/NatClass.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ex/NatClass.ML Fri Jan 03 15:01:55 1997 +0100
@@ -1,7 +1,6 @@
(* Title: FOL/ex/NatClass.ML
ID: $Id$
- Author: Lawrence C Paulson, Cambridge University Computer Laboratory
- Copyright 1992 University of Cambridge
+ Author: Markus Wenzel, TU Muenchen
This is Nat.ML trivially modified to make it work with NatClass.thy.
*)
@@ -36,27 +35,27 @@
by (rtac rec_Suc 1);
qed "add_Suc";
-val add_ss = FOL_ss addsimps [add_0, add_Suc];
+Addsimps [add_0, add_Suc];
goal NatClass.thy "(k+m)+n = k+(m+n)";
by (res_inst_tac [("n","k")] induct 1);
-by (simp_tac add_ss 1);
-by (asm_simp_tac add_ss 1);
+by (Simp_tac 1);
+by (Asm_simp_tac 1);
qed "add_assoc";
goal NatClass.thy "m+0 = m";
by (res_inst_tac [("n","m")] induct 1);
-by (simp_tac add_ss 1);
-by (asm_simp_tac add_ss 1);
+by (Simp_tac 1);
+by (Asm_simp_tac 1);
qed "add_0_right";
goal NatClass.thy "m+Suc(n) = Suc(m+n)";
by (res_inst_tac [("n","m")] induct 1);
-by (ALLGOALS (asm_simp_tac add_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "add_Suc_right";
val [prem] = goal NatClass.thy "(!!n. f(Suc(n)) = Suc(f(n))) ==> f(i+j) = i+f(j)";
by (res_inst_tac [("n","i")] induct 1);
-by (simp_tac add_ss 1);
-by (asm_simp_tac (add_ss addsimps [prem]) 1);
+by (Simp_tac 1);
+by (asm_simp_tac (!simpset addsimps [prem]) 1);
result();
--- a/src/FOL/ex/cla.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ex/cla.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,17 +11,17 @@
open Cla; (*in case structure Int is open!*)
goal FOL.thy "(P --> Q | R) --> (P-->Q) | (P-->R)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*If and only if*)
goal FOL.thy "(P<->Q) <-> (Q<->P)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
goal FOL.thy "~ (P <-> ~P)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
@@ -38,132 +38,132 @@
writeln"Pelletier's examples";
(*1*)
goal FOL.thy "(P-->Q) <-> (~Q --> ~P)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*2*)
goal FOL.thy "~ ~ P <-> P";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*3*)
goal FOL.thy "~(P-->Q) --> (Q-->P)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*4*)
goal FOL.thy "(~P-->Q) <-> (~Q --> P)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*5*)
goal FOL.thy "((P|Q)-->(P|R)) --> (P|(Q-->R))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*6*)
goal FOL.thy "P | ~ P";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*7*)
goal FOL.thy "P | ~ ~ ~ P";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*8. Peirce's law*)
goal FOL.thy "((P-->Q) --> P) --> P";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*9*)
goal FOL.thy "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*10*)
goal FOL.thy "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P<->Q)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*11. Proved in each direction (incorrectly, says Pelletier!!) *)
goal FOL.thy "P<->P";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*12. "Dijkstra's law"*)
goal FOL.thy "((P <-> Q) <-> R) <-> (P <-> (Q <-> R))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*13. Distributive law*)
goal FOL.thy "P | (Q & R) <-> (P | Q) & (P | R)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*14*)
goal FOL.thy "(P <-> Q) <-> ((Q | ~P) & (~Q|P))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*15*)
goal FOL.thy "(P --> Q) <-> (~P | Q)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*16*)
goal FOL.thy "(P-->Q) | (Q-->P)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*17*)
goal FOL.thy "((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Classical Logic: examples with quantifiers";
goal FOL.thy "(ALL x. P(x) & Q(x)) <-> (ALL x. P(x)) & (ALL x. Q(x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
goal FOL.thy "(EX x. P-->Q(x)) <-> (P --> (EX x.Q(x)))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
goal FOL.thy "(EX x.P(x)-->Q) <-> (ALL x.P(x)) --> Q";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
goal FOL.thy "(ALL x.P(x)) | Q <-> (ALL x. P(x) | Q)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*Discussed in Avron, Gentzen-Type Systems, Resolution and Tableaux,
JAR 10 (265-281), 1993. Proof is trivial!*)
goal FOL.thy
"~ ((EX x.~P(x)) & ((EX x.P(x)) | (EX x.P(x) & Q(x))) & ~ (EX x.P(x)))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problems requiring quantifier duplication";
(*Needs multiple instantiation of ALL.*)
goal FOL.thy "(ALL x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
(*Needs double instantiation of the quantifier*)
goal FOL.thy "EX x. P(x) --> P(a) & P(b)";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
goal FOL.thy "EX z. P(z) --> (ALL x. P(x))";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
goal FOL.thy "EX x. (EX y. P(y)) --> P(x)";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
(*V. Lifschitz, What Is the Inverse Method?, JAR 5 (1989), 1--23. NOT PROVED*)
@@ -176,40 +176,40 @@
writeln"Problem 18";
goal FOL.thy "EX y. ALL x. P(y)-->P(x)";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
writeln"Problem 19";
goal FOL.thy "EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
writeln"Problem 20";
goal FOL.thy "(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 (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 21";
goal FOL.thy "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> (EX x. P<->Q(x))";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
writeln"Problem 22";
goal FOL.thy "(ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 23";
goal FOL.thy "(ALL x. P | Q(x)) <-> (P | (ALL x. Q(x)))";
-by (best_tac FOL_cs 1);
+by (Best_tac 1);
result();
writeln"Problem 24";
goal FOL.thy "~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) & \
\ ~(EX x.P(x)) --> (EX x.Q(x)) & (ALL x. Q(x)|R(x) --> S(x)) \
\ --> (EX x. P(x)&R(x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 25";
@@ -218,14 +218,14 @@
\ (ALL x. P(x) --> (M(x) & L(x))) & \
\ ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x))) \
\ --> (EX x. Q(x)&P(x))";
-by (best_tac FOL_cs 1);
+by (Best_tac 1);
result();
writeln"Problem 26";
goal FOL.thy "((EX x. p(x)) <-> (EX x. q(x))) & \
\ (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y))) \
\ --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 27";
@@ -234,7 +234,7 @@
\ (ALL x. M(x) & L(x) --> P(x)) & \
\ ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x))) \
\ --> (ALL x. M(x) --> ~L(x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 28. AMENDED";
@@ -242,21 +242,21 @@
\ ((ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) & \
\ ((EX x.S(x)) --> (ALL x. L(x) --> M(x))) \
\ --> (ALL x. P(x) & L(x) --> M(x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 29. Essentially the same as Principia Mathematica *11.71";
goal FOL.thy "(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 (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 30";
goal FOL.thy "(ALL x. P(x) | Q(x) --> ~ R(x)) & \
\ (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x)) \
\ --> (ALL x. S(x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 31";
@@ -264,7 +264,7 @@
\ (EX x. L(x) & P(x)) & \
\ (ALL x. ~ R(x) --> M(x)) \
\ --> (EX x. L(x) & M(x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 32";
@@ -272,13 +272,13 @@
\ (ALL x. S(x) & R(x) --> L(x)) & \
\ (ALL x. M(x) --> R(x)) \
\ --> (ALL x. P(x) & M(x) --> L(x))";
-by (best_tac FOL_cs 1);
+by (Best_tac 1);
result();
writeln"Problem 33";
goal FOL.thy "(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 (best_tac FOL_cs 1);
+by (Best_tac 1);
result();
writeln"Problem 34 AMENDED (TWICE!!)";
@@ -287,13 +287,13 @@
\ ((EX x. q(x)) <-> (ALL y. p(y)))) <-> \
\ ((EX x. ALL y. q(x) <-> q(y)) <-> \
\ ((EX x. p(x)) <-> (ALL y. q(y))))";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
writeln"Problem 35";
goal FOL.thy "EX x y. P(x,y) --> (ALL u v. P(u,v))";
by (mini_tac 1);
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
(*Without miniscope, would have to use deepen_tac; would be slower*)
result();
@@ -303,7 +303,7 @@
\ (ALL x. EX y. G(x,y)) & \
\ (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 (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 37";
@@ -312,7 +312,7 @@
\ (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) & \
\ ((EX x y. Q(x,y)) --> (ALL x. R(x,x))) \
\ --> (ALL x. EX y. R(x,y))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 38";
@@ -322,40 +322,39 @@
\ (ALL x. (~p(a) | p(x) | (EX z. EX w. p(z) & r(x,w) & r(w,z))) & \
\ (~p(a) | ~(EX y. p(y) & r(x,y)) | \
\ (EX z. EX w. p(z) & r(x,w) & r(w,z))))";
-by (deepen_tac FOL_cs 0 1); (*beats fast_tac!*)
+by (Deepen_tac 0 1); (*beats fast_tac!*)
result();
writeln"Problem 39";
goal FOL.thy "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 40. AMENDED";
goal FOL.thy "(EX y. ALL x. F(x,y) <-> F(x,x)) --> \
\ ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 41";
goal FOL.thy "(ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x)) \
\ --> ~ (EX z. ALL x. f(x,z))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 42";
goal FOL.thy "~ (EX y. ALL x. p(x,y) <-> ~ (EX z. p(x,z) & p(z,x)))";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
writeln"Problem 43";
goal FOL.thy "(ALL x. ALL y. q(x,y) <-> (ALL z. p(z,x) <-> p(z,y))) \
\ --> (ALL x. ALL y. q(x,y) <-> q(y,x))";
-by (mini_tac 1);
-by (deepen_tac FOL_cs 5 1);
-(*Faster alternative proof!
- by (asm_simp_tac FOL_ss 1); by (fast_tac FOL_cs 1);
+by (Auto_tac());
+(*The proof above cheats by using rewriting! A purely logical proof is
+ by (mini_tac 1 THEN Deepen_tac 5 1);
Can use just deepen_tac but it requires 253 secs?!?
- by (deepen_tac FOL_cs 0 1);
+ by (Deepen_tac 0 1);
*)
result();
@@ -364,7 +363,7 @@
\ (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) & \
\ (EX x. j(x) & (ALL y. g(y) --> h(x,y))) \
\ --> (EX x. j(x) & ~f(x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 45";
@@ -374,7 +373,7 @@
\ (EX x. f(x) & (ALL y. h(x,y) --> l(y)) \
\ & (ALL y. g(y) & h(x,y) --> j(x,y))) \
\ --> (EX x. f(x) & ~ (EX y. g(y) & h(x,y)))";
-by (best_tac FOL_cs 1);
+by (Best_tac 1);
result();
@@ -382,7 +381,7 @@
writeln"Problem 48";
goal FOL.thy "(a=b | c=d) & (a=c | b=d) --> a=d | b=c";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 49 NOT PROVED AUTOMATICALLY";
@@ -390,26 +389,26 @@
the type constraint ensures that x,y,z have the same type as a,b,u. *)
goal FOL.thy "(EX x y::'a. ALL z. z=x | z=y) & P(a) & P(b) & a~=b \
\ --> (ALL u::'a.P(u))";
-by (safe_tac FOL_cs);
+by (Step_tac 1);
by (res_inst_tac [("x","a")] allE 1);
by (assume_tac 1);
by (res_inst_tac [("x","b")] allE 1);
by (assume_tac 1);
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 50";
(*What has this to do with equality?*)
goal FOL.thy "(ALL x. P(a,x) | (ALL y.P(x,y))) --> (EX x. ALL y.P(x,y))";
by (mini_tac 1);
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
writeln"Problem 51";
goal FOL.thy
"(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 (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 52";
@@ -417,7 +416,7 @@
goal FOL.thy
"(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 (best_tac FOL_cs 1);
+by (Best_tac 1);
result();
writeln"Problem 55";
@@ -436,12 +435,12 @@
\ (ALL x. EX y. ~hates(x,y)) & \
\ ~ agatha=butler --> \
\ killed(?who,agatha)";
-by (safe_tac FOL_cs);
+by (Step_tac 1);
by (dres_inst_tac [("x1","x")] (spec RS mp) 1);
by (assume_tac 1);
by (etac (spec RS exE) 1);
by (REPEAT (etac allE 1));
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
****)
@@ -456,37 +455,37 @@
\ (ALL x. hates(agatha,x) --> hates(butler,x)) & \
\ (ALL x. ~hates(x,agatha) | ~hates(x,butler) | ~hates(x,charles)) --> \
\ killed(?who,agatha)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 56";
goal FOL.thy
"(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 57";
goal FOL.thy
"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 (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 58 NOT PROVED AUTOMATICALLY";
goal FOL.thy "(ALL x y. f(x)=g(y)) --> (ALL x y. f(f(x))=f(g(y)))";
-by (slow_tac (FOL_cs addEs [subst_context]) 1);
+by (slow_tac (!claset addEs [subst_context]) 1);
result();
writeln"Problem 59";
goal FOL.thy "(ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))";
-by (deepen_tac FOL_cs 0 1);
+by (Deepen_tac 0 1);
result();
writeln"Problem 60";
goal FOL.thy
"ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
writeln"Problem 62 as corrected in AAR newletter #31";
@@ -494,7 +493,7 @@
"(ALL x. p(a) & (p(x) --> p(f(x))) --> p(f(f(x)))) <-> \
\ (ALL x. (~p(a) | p(x) | p(f(f(x)))) & \
\ (~p(a) | ~p(f(x)) | p(f(f(x)))))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
(*Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
--- a/src/FOL/ex/intro.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/ex/intro.ML Fri Jan 03 15:01:55 1997 +0100
@@ -42,16 +42,16 @@
result();
-(**** Demonstration of fast_tac ****)
+(**** Demonstration of Fast_tac ****)
goal FOL.thy "(EX y. ALL x. J(y,x) <-> ~J(x,x)) \
\ --> ~ (ALL x. EX y. ALL z. J(z,y) <-> ~ J(z,x))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
goal FOL.thy "ALL x. P(x,f(x)) <-> \
\ (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
result();
--- a/src/FOL/simpdata.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/FOL/simpdata.ML Fri Jan 03 15:01:55 1997 +0100
@@ -187,10 +187,33 @@
open Simplifier Induction;
+(*** Integration of simplifier with classical reasoner ***)
+
+(*Add a simpset to a classical set!*)
+infix 4 addss;
+fun cs addss ss = cs addbefore asm_full_simp_tac ss 1;
+
+fun Addss ss = (claset := !claset addbefore asm_full_simp_tac ss 1);
+
+(*Designed to be idempotent, except if best_tac instantiates variables
+ in some of the subgoals*)
+fun auto_tac (cs,ss) =
+ ALLGOALS (asm_full_simp_tac ss) THEN
+ REPEAT (safe_tac cs THEN ALLGOALS (asm_full_simp_tac ss)) THEN
+ REPEAT (FIRSTGOAL (best_tac (cs addss ss))) THEN
+ prune_params_tac;
+
+fun Auto_tac() = auto_tac (!claset, !simpset);
+
+fun auto() = by (Auto_tac());
+
+
(*Add congruence rules for = or <-> (instead of ==) *)
infix 4 addcongs;
fun ss addcongs congs =
- ss addeqcongs (congs RL [eq_reflection,iff_reflection]);
+ ss addeqcongs (congs RL [eq_reflection,iff_reflection]);
+
+fun Addcongs congs = (simpset := !simpset addcongs congs);
(*Add a simpset to a classical set!*)
infix 4 addss;
@@ -223,3 +246,24 @@
val FOL_ss = IFOL_ss addsimps (cla_simps @ ex_simps @ all_simps);
+
+
+(*** Install simpsets and datatypes in theory structure ***)
+
+simpset := FOL_ss;
+
+exception SS_DATA of simpset;
+
+let fun merge [] = SS_DATA empty_ss
+ | merge ss = let val ss = map (fn SS_DATA x => x) ss;
+ in SS_DATA (foldl merge_ss (hd ss, tl ss)) end;
+
+ fun put (SS_DATA ss) = simpset := ss;
+
+ fun get () = SS_DATA (!simpset);
+in add_thydata "FOL"
+ ("simpset", ThyMethods {merge = merge, put = put, get = get})
+end;
+
+
+add_thy_reader_file "thy_data.ML";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/FOL/thy_data.ML Fri Jan 03 15:01:55 1997 +0100
@@ -0,0 +1,19 @@
+(* Title: FOL/thy_data.ML
+ ID: $Id$
+ Author: Carsten Clasohm
+ Copyright 1995 TU Muenchen
+
+Definitions that have to be reread after init_thy_reader has been invoked
+*)
+
+fun simpset_of tname =
+ case get_thydata tname "simpset" of
+ None => empty_ss
+ | Some (SS_DATA ss) => ss;
+
+fun claset_of tname =
+ case get_thydata tname "claset" of
+ None => empty_cs
+ | Some (CS_DATA cs) => cs;
+
+
--- a/src/ZF/AC.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC.ML Fri Jan 03 15:01:55 1997 +0100
@@ -12,7 +12,7 @@
val [nonempty] = goal AC.thy
"[| !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)";
by (excluded_middle_tac "A=0" 1);
-by (asm_simp_tac (ZF_ss addsimps [Pi_empty1]) 2 THEN fast_tac ZF_cs 2);
+by (asm_simp_tac (!simpset addsimps [Pi_empty1]) 2 THEN Fast_tac 2);
(*The non-trivial case*)
by (fast_tac (eq_cs addIs [AC, nonempty]) 1);
qed "AC_Pi";
@@ -35,7 +35,7 @@
\ |] ==> EX f: A->Union(A). ALL x:A. f`x : x";
by (res_inst_tac [("B1", "%x.x")] (AC_Pi RS exE) 1);
by (etac nonempty 1);
-by (fast_tac (ZF_cs addDs [apply_type] addIs [Pi_type]) 1);
+by (fast_tac (!claset addDs [apply_type] addIs [Pi_type]) 1);
qed "AC_func";
goal ZF.thy "!!x A. [| 0 ~: A; x: A |] ==> EX y. y:x";
@@ -53,7 +53,7 @@
by (rtac bexI 2);
by (assume_tac 2);
by (etac fun_weaken_type 2);
-by (ALLGOALS (fast_tac ZF_cs));
+by (ALLGOALS (Fast_tac));
qed "AC_func_Pow";
goal AC.thy "!!A. 0 ~: A ==> EX f. f: (PROD x:A. x)";
--- a/src/ZF/AC.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC.thy Fri Jan 03 15:01:55 1997 +0100
@@ -8,7 +8,7 @@
This definition comes from Halmos (1960), page 59.
*)
-AC = ZF + "func" +
+AC = func +
rules
AC "[| a: A; !!x. x:A ==> (EX y. y:B(x)) |] ==> EX z. z : Pi(A,B)"
end
--- a/src/ZF/AC/AC0_AC1.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC0_AC1.ML Fri Jan 03 15:01:55 1997 +0100
@@ -7,11 +7,11 @@
*)
goal thy "!!A. 0~:A ==> A <= Pow(Union(A))-{0}";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val subset_Pow_Union = result();
goal thy "!!f. [| f:(PROD X:A. X); D<=A |] ==> EX g. g:(PROD X:D. X)";
-by (fast_tac (AC_cs addSIs [restrict_type, apply_type]) 1);
+by (fast_tac (!claset addSIs [restrict_type, apply_type]) 1);
val lemma1 = result();
goalw thy AC_defs "!!Z. AC0 ==> AC1";
@@ -19,8 +19,8 @@
qed "AC0_AC1";
goalw thy AC_defs "!!Z. AC1 ==> AC0";
-by (deepen_tac ZF_cs 0 1);
+by (Deepen_tac 0 1);
(*Large search space. Faster proof by
- by (fast_tac (FOL_cs addSIs [notI, singletonI] addSEs [notE, DiffE]) 1);
+ by (fast_tac (!claset addSIs [notI, singletonI] addSEs [notE, DiffE]) 1);
*)
qed "AC1_AC0";
--- a/src/ZF/AC/AC10_AC15.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC10_AC15.ML Fri Jan 03 15:01:55 1997 +0100
@@ -30,7 +30,7 @@
goalw thy [lepoll_def] "!!A. A~=0 ==> B lepoll A*B";
by (etac not_emptyE 1);
by (res_inst_tac [("x","lam z:B. <x,z>")] exI 1);
-by (fast_tac (AC_cs addSIs [snd_conv, lam_injective]) 1);
+by (fast_tac (!claset addSIs [snd_conv, lam_injective]) 1);
val lepoll_Sigma = result();
goal thy "!!A. 0~:A ==> ALL B:{cons(0,x*nat). x:A}. ~Finite(B)";
@@ -41,18 +41,18 @@
by (dresolve_tac [subset_consI RS subset_imp_lepoll RS lepoll_Finite] 1);
by (resolve_tac [lepoll_Sigma RS lepoll_Finite RS (nat_not_Finite RS notE)] 1
THEN (assume_tac 2));
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val cons_times_nat_not_Finite = result();
goal thy "!!A. [| Union(C)=A; a:A |] ==> EX B:C. a:B & B <= A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val lemma1 = result();
goalw thy [pairwise_disjoint_def]
"!!A. [| pairwise_disjoint(A); B:A; C:A; a:B; a:C |] ==> B=C";
by (dtac IntI 1 THEN (assume_tac 1));
by (dres_inst_tac [("A","B Int C")] not_emptyI 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val lemma2 = result();
goalw thy [sets_of_size_between_def]
@@ -62,12 +62,12 @@
\ 0:u & 2 lepoll u & u lepoll n";
by (rtac ballI 1);
by (etac ballE 1);
-by (fast_tac ZF_cs 2);
+by (Fast_tac 2);
by (REPEAT (etac conjE 1));
by (dresolve_tac [consI1 RSN (2, lemma1)] 1);
by (etac bexE 1);
by (rtac ex1I 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (REPEAT (etac conjE 1));
by (rtac lemma2 1 THEN (REPEAT (assume_tac 1)));
val lemma3 = result();
@@ -79,13 +79,13 @@
by (res_inst_tac [("d", "%y. P(converse(f)`y)")] lam_injective 1);
by (etac RepFunE 1);
by (forward_tac [inj_is_fun RS apply_type] 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addIs [LeastI2]
+by (fast_tac (!claset addIs [LeastI2]
addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1);
by (etac RepFunE 1);
by (rtac LeastI2 1);
-by (fast_tac AC_cs 1);
-by (fast_tac (AC_cs addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1);
-by (fast_tac (AC_cs addEs [sym, left_inverse RS ssubst]) 1);
+by (Fast_tac 1);
+by (fast_tac (!claset addSEs [Ord_in_Ord, inj_is_fun RS apply_type]) 1);
+by (fast_tac (!claset addEs [sym, left_inverse RS ssubst]) 1);
val lemma4 = result();
goal thy "!!A. [| n:nat; B:A; u(B) <= cons(0, B*nat); 0:u(B); 2 lepoll u(B); \
@@ -93,15 +93,15 @@
\ ==> (lam x:A. {fst(x). x:u(x)-{0}})`B ~= 0 & \
\ (lam x:A. {fst(x). x:u(x)-{0}})`B <= B & \
\ (lam x:A. {fst(x). x:u(x)-{0}})`B lepoll n";
-by (asm_simp_tac AC_ss 1);
+by (Asm_simp_tac 1);
by (rtac conjI 1);
by (fast_tac (empty_cs addSDs [RepFun_eq_0_iff RS iffD1]
addDs [lepoll_Diff_sing]
addEs [lepoll_trans RS succ_lepoll_natE, ssubst]
addSIs [notI, lepoll_refl, nat_0I]) 1);
by (rtac conjI 1);
-by (fast_tac (ZF_cs addSIs [fst_type] addSEs [consE]) 1);
-by (fast_tac (ZF_cs addSEs [equalityE,
+by (fast_tac (!claset addSIs [fst_type] addSEs [consE]) 1);
+by (fast_tac (!claset addSEs [equalityE,
Diff_lepoll RS (nat_into_Ord RSN (2, lemma4))]) 1);
val lemma5 = result();
@@ -125,7 +125,7 @@
goalw thy AC_defs "!!Z. [| n:nat; 1 le n; AC10(n) |] ==> AC11";
by (rtac bexI 1 THEN (assume_tac 2));
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "AC10_AC11";
(* ********************************************************************** *)
@@ -141,11 +141,11 @@
(* ********************************************************************** *)
goalw thy AC_defs "!!Z. AC12 ==> AC15";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (etac allE 1);
by (etac impE 1);
by (etac cons_times_nat_not_Finite 1);
-by (fast_tac (ZF_cs addSIs [ex_fun_AC13_AC15]) 1);
+by (fast_tac (!claset addSIs [ex_fun_AC13_AC15]) 1);
qed "AC12_AC15";
(* ********************************************************************** *)
@@ -167,7 +167,7 @@
(* ********************************************************************** *)
goalw thy AC_defs "!!n. [| n:nat; 1 le n; AC10(n) |] ==> AC13(n)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (fast_tac (empty_cs addSEs [allE, cons_times_nat_not_Finite RSN (2, impE),
ex_fun_AC13_AC15]) 1);
qed "AC10_AC13";
@@ -187,10 +187,10 @@
by (mp_tac 1);
by (etac exE 1);
by (res_inst_tac [("x","lam x:A. {f`x}")] exI 1);
-by (asm_full_simp_tac (AC_ss addsimps
+by (asm_full_simp_tac (!simpset addsimps
[singleton_eqpoll_1 RS eqpoll_imp_lepoll,
singletonI RS not_emptyI]) 1);
-by (fast_tac (AC_cs addSEs [apply_type]) 1);
+by (fast_tac (!claset addSEs [apply_type]) 1);
qed "AC1_AC13";
(* ********************************************************************** *)
@@ -199,7 +199,7 @@
goalw thy AC_defs "!!m n. [| m:nat; n:nat; m le n; AC13(m) |] ==> AC13(n)";
by (dtac nat_le_imp_lepoll 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (ZF_cs addSEs [lepoll_trans]) 1);
+by (fast_tac (!claset addSEs [lepoll_trans]) 1);
qed "AC13_mono";
(* ********************************************************************** *)
@@ -219,7 +219,7 @@
(* ********************************************************************** *)
goalw thy AC_defs "!!Z. AC14 ==> AC15";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "AC14_AC15";
(* ********************************************************************** *)
@@ -231,7 +231,7 @@
(* ********************************************************************** *)
goal thy "!!A. [| A~=0; A lepoll 1 |] ==> EX a. A={a}";
-by (fast_tac (AC_cs addSEs [not_emptyE, lepoll_1_is_sing]) 1);
+by (fast_tac (!claset addSEs [not_emptyE, lepoll_1_is_sing]) 1);
val lemma_aux = result();
goal thy "!!f. ALL B:A. f(B)~=0 & f(B)<=B & f(B) lepoll 1 \
@@ -240,12 +240,12 @@
by (dtac bspec 1 THEN (assume_tac 1));
by (REPEAT (etac conjE 1));
by (eresolve_tac [lemma_aux RS exE] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (AC_ss addsimps [the_element]) 1);
-by (fast_tac (AC_cs addEs [ssubst]) 1);
+by (asm_full_simp_tac (!simpset addsimps [the_element]) 1);
+by (fast_tac (!claset addEs [ssubst]) 1);
val lemma = result();
goalw thy AC_defs "!!Z. AC13(1) ==> AC1";
-by (fast_tac (AC_cs addSEs [lemma]) 1);
+by (fast_tac (!claset addSEs [lemma]) 1);
qed "AC13_AC1";
(* ********************************************************************** *)
@@ -253,5 +253,5 @@
(* ********************************************************************** *)
goalw thy [AC11_def, AC14_def] "!!Z. AC11 ==> AC14";
-by (fast_tac (ZF_cs addSIs [AC10_AC13]) 1);
+by (fast_tac (!claset addSIs [AC10_AC13]) 1);
qed "AC11_AC14";
--- a/src/ZF/AC/AC15_WO6.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC15_WO6.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,15 +8,15 @@
open AC15_WO6;
goal thy "!!x. Ord(x) ==> (UN a<x. F(a)) = (UN a:x. F(a))";
-by (fast_tac (AC_cs addSIs [equalityI, ltI] addSDs [ltD]) 1);
+by (fast_tac (!claset addSIs [equalityI, ltI] addSDs [ltD]) 1);
val OUN_eq_UN = result();
val [prem] = goal thy "ALL x:Pow(A)-{0}. f`x~=0 & f`x<=x & f`x lepoll m ==> \
\ (UN i<LEAST x. HH(f,A,x)={A}. HH(f,A,i)) = A";
-by (simp_tac (AC_ss addsimps [Ord_Least RS OUN_eq_UN]) 1);
+by (simp_tac (!simpset addsimps [Ord_Least RS OUN_eq_UN]) 1);
by (rtac equalityI 1);
-by (fast_tac (AC_cs addSDs [less_Least_subset_x]) 1);
-by (fast_tac (AC_cs addSDs [prem RS bspec]
+by (fast_tac (!claset addSDs [less_Least_subset_x]) 1);
+by (fast_tac (!claset addSDs [prem RS bspec]
addSIs [f_subsets_imp_UN_HH_eq_x RS (Diff_eq_0_iff RS iffD1)]) 1);
val lemma1 = result();
@@ -26,7 +26,7 @@
by (dresolve_tac [ltD RS less_Least_subset_x] 1);
by (forward_tac [HH_subset_imp_eq] 1);
by (etac ssubst 1);
-by (fast_tac (AC_cs addIs [prem RS ballE]
+by (fast_tac (!claset addIs [prem RS ballE]
addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1);
val lemma2 = result();
@@ -34,13 +34,13 @@
by (rtac allI 1);
by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
by (etac impE 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (REPEAT (eresolve_tac [bexE,conjE,exE] 1));
by (rtac bexI 1 THEN (assume_tac 2));
by (rtac conjI 1 THEN (assume_tac 1));
by (res_inst_tac [("x","LEAST i. HH(f,A,i)={A}")] exI 1);
by (res_inst_tac [("x","lam j: (LEAST i. HH(f,A,i)={A}). HH(f,A,j)")] exI 1);
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSIs [Ord_Least, lam_type RS domain_of_fun]
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSIs [Ord_Least, lam_type RS domain_of_fun]
addSEs [less_Least_subset_x, lemma1, lemma2]) 1);
qed "AC15_WO6";
--- a/src/ZF/AC/AC16_WO4.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC16_WO4.ML Fri Jan 03 15:01:55 1997 +0100
@@ -19,9 +19,9 @@
by (etac exE 1);
by (res_inst_tac [("x","n")] exI 1);
by (res_inst_tac [("x","lam i:n. {f`i}")] exI 1);
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
by (rewrite_goals_tac [bij_def, surj_def]);
-by (fast_tac (AC_cs addSIs [ltI, nat_into_Ord, lam_funtype RS domain_of_fun,
+by (fast_tac (!claset addSIs [ltI, nat_into_Ord, lam_funtype RS domain_of_fun,
equalityI, singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
nat_1_lepoll_iff RS iffD2]
addSEs [apply_type, ltE]) 1);
@@ -35,12 +35,12 @@
val well_ord_paired = standard (paired_bij RS bij_is_inj RS well_ord_rvimage);
goal thy "!!A. [| A lepoll B; ~ A lepoll C |] ==> ~ B lepoll C";
-by (fast_tac (FOL_cs addEs [notE, lepoll_trans]) 1);
+by (fast_tac (!claset addEs [notE, lepoll_trans]) 1);
val lepoll_trans1 = result();
goalw thy [lepoll_def]
"!!X.[| Y lepoll X; well_ord(X, R) |] ==> EX S. well_ord(Y, S)";
-by (fast_tac (AC_cs addSEs [well_ord_rvimage]) 1);
+by (fast_tac (!claset addSEs [well_ord_rvimage]) 1);
val well_ord_lepoll = result();
goal thy "!!X. [| well_ord(X,R); well_ord(Y,S) \
@@ -57,7 +57,7 @@
by (res_inst_tac [("x","{{a,x}. a:nat Un Hartog(z)}")] exI 1);
by (resolve_tac [Ord_nat RS well_ord_Memrel RS (Ord_Hartog RS
well_ord_Memrel RSN (2, well_ord_Un)) RS exE] 1);
-by (fast_tac (AC_cs addSIs [Ord_Hartog, well_ord_Memrel, well_ord_paired,
+by (fast_tac (!claset addSIs [Ord_Hartog, well_ord_Memrel, well_ord_paired,
equals0I, HartogI RSN (2, lepoll_trans1),
subset_imp_lepoll RS (paired_eqpoll RS eqpoll_sym RS
eqpoll_imp_lepoll RSN (2, lepoll_trans))]
@@ -68,7 +68,7 @@
val lemma2 = result();
val [prem] = goal thy "~Finite(B) ==> ~Finite(A Un B)";
-by (fast_tac (AC_cs
+by (fast_tac (!claset
addSIs [subset_imp_lepoll RS (prem RSN (2, lepoll_infinite))]) 1);
val infinite_Un = result();
@@ -90,13 +90,13 @@
by (res_inst_tac [("d","%z. if(z=y, A, converse(f)`z)")] lam_injective 1);
by (ALLGOALS
(asm_simp_tac
- (ZF_ss addsimps [inj_is_fun RS apply_type, left_inverse]
- setloop (split_tac [expand_if] ORELSE' step_tac ZF_cs))));
+ (!simpset addsimps [inj_is_fun RS apply_type, left_inverse]
+ setloop (split_tac [expand_if] ORELSE' Step_tac))));
val succ_not_lepoll_lemma = result();
goalw thy [lepoll_def, eqpoll_def, bij_def, surj_def]
"!!A. [| ~A eqpoll B; A lepoll B |] ==> succ(A) lepoll B";
-by (fast_tac (AC_cs addSEs [succ_not_lepoll_lemma, inj_is_fun]) 1);
+by (fast_tac (!claset addSEs [succ_not_lepoll_lemma, inj_is_fun]) 1);
val succ_not_lepoll_imp_eqpoll = result();
val [prem] = goalw thy [s_u_def]
@@ -108,7 +108,7 @@
by (etac CollectE 1);
by (etac conjE 1);
by (etac swap 1);
-by (fast_tac (AC_cs addSEs [succ_not_lepoll_imp_eqpoll]) 1);
+by (fast_tac (!claset addSEs [succ_not_lepoll_imp_eqpoll]) 1);
val suppose_not = result();
(* ********************************************************************** *)
@@ -130,13 +130,13 @@
by (etac nat_lepoll_imp_ex_eqpoll_n 1);
by (resolve_tac [ordertype_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll
RSN (2, lepoll_trans)] 1 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSIs [nat_le_infinite_Ord RS le_imp_lepoll]
+by (fast_tac (!claset addSIs [nat_le_infinite_Ord RS le_imp_lepoll]
addSEs [Ord_ordertype, ordertype_eqpoll RS eqpoll_imp_lepoll
RS lepoll_infinite]) 1);
val ex_subset_eqpoll_n = result();
goalw thy [lesspoll_def] "!!n. n: nat ==> n lesspoll nat";
-by (fast_tac (AC_cs addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_lepoll,
+by (fast_tac (!claset addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_lepoll,
eqpoll_sym RS eqpoll_imp_lepoll]
addIs [Ord_nat RSN (2, nat_succI RS ltI) RS leI
RS le_imp_lepoll RS lepoll_trans RS succ_lepoll_natE]) 1);
@@ -157,27 +157,27 @@
val Diff_Finite_eqpoll = result();
goal thy "!!x. [| a<=y; b:y-a; u:x |] ==> cons(b, cons(u, a)) : Pow(x Un y)";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val cons_cons_subset = result();
goal thy "!!x. [| a eqpoll k; a<=y; b:y-a; u:x; x Int y = 0 \
\ |] ==> cons(b, cons(u, a)) eqpoll succ(succ(k))";
-by (fast_tac (AC_cs addSIs [cons_eqpoll_succ] addEs [equals0D]) 1);
+by (fast_tac (!claset addSIs [cons_eqpoll_succ] addEs [equals0D]) 1);
val cons_cons_eqpoll = result();
goalw thy [s_u_def] "s_u(u, t_n, k, y) <= t_n";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val s_u_subset = result();
goalw thy [s_u_def, succ_def]
"!!w. [| w:t_n; cons(b,cons(u,a)) <= w; a <= y; b : y-a; k eqpoll a \
\ |] ==> w: s_u(u, t_n, succ(k), y)";
-by (fast_tac (AC_cs addDs [eqpoll_imp_lepoll RS cons_lepoll_cong]
+by (fast_tac (!claset addDs [eqpoll_imp_lepoll RS cons_lepoll_cong]
addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
val s_uI = result();
goalw thy [s_u_def] "!!v. v : s_u(u, t_n, k, y) ==> u : v";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val in_s_u_imp_u_in = result();
goal thy
@@ -187,33 +187,33 @@
\ ==> EX! c. c:{v:s_u(u, t_n, succ(k), y). a <= v} & b:c";
by (etac ballE 1);
by (fast_tac (FOL_cs addSIs [CollectI, cons_cons_subset,
- eqpoll_sym RS cons_cons_eqpoll]) 2);
+ eqpoll_sym RS cons_cons_eqpoll]) 2);
by (etac ex1E 1);
by (res_inst_tac [("a","w")] ex1I 1);
by (rtac conjI 1);
by (rtac CollectI 1);
by (fast_tac (FOL_cs addSEs [s_uI]) 1);
-by (fast_tac AC_cs 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
+by (Fast_tac 1);
by (etac allE 1);
by (etac impE 1);
by (assume_tac 2);
-by (fast_tac (AC_cs addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
+by (fast_tac (!claset addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
val ex1_superset_a = result();
goal thy
"!!A. [| succ(k) eqpoll A; k eqpoll B; B <= A; a : A-B; k:nat \
\ |] ==> A = cons(a, B)";
by (rtac equalityI 1);
-by (fast_tac AC_cs 2);
+by (Fast_tac 2);
by (resolve_tac [Diff_eq_0_iff RS iffD1] 1);
by (rtac equals0I 1);
by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll] 1);
by (dresolve_tac [eqpoll_sym RS cons_eqpoll_succ] 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (dtac cons_eqpoll_succ 1);
-by (fast_tac AC_cs 1);
-by (fast_tac (AC_cs addSIs [nat_succI]
+by (Fast_tac 1);
+by (fast_tac (!claset addSIs [nat_succI]
addSEs [[eqpoll_sym RS eqpoll_imp_lepoll, subset_imp_lepoll] MRS
(lepoll_trans RS lepoll_trans) RS succ_lepoll_natE]) 1);
val set_eq_cons = result();
@@ -227,19 +227,19 @@
\ Int y = cons(b, a)";
by (dresolve_tac [ex1_superset_a RS theI] 1 THEN REPEAT (assume_tac 1));
by (rtac set_eq_cons 1);
-by (REPEAT (fast_tac AC_cs 1));
+by (REPEAT (Fast_tac 1));
val the_eq_cons = result();
goal thy "!!a. [| cons(x,a) = cons(y,a); x~: a |] ==> x = y ";
-by (fast_tac (AC_cs addSEs [equalityE]) 1);
+by (fast_tac (!claset addSEs [equalityE]) 1);
val cons_eqE = result();
goal thy "!!A. A = B ==> A Int C = B Int C";
-by (asm_simp_tac AC_ss 1);
+by (Asm_simp_tac 1);
val eq_imp_Int_eq = result();
goal thy "!!a. [| a=b; a=c; b=d |] ==> c=d";
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val msubst = result();
(* ********************************************************************** *)
@@ -271,7 +271,7 @@
by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
by (rtac impI 1);
by (rtac cons_eqE 1);
-by (fast_tac AC_cs 2);
+by (Fast_tac 2);
by (dres_inst_tac [("A","THE c. ?P(c)"), ("C","y")] eq_imp_Int_eq 1);
by (eresolve_tac [[asm_rl, the_eq_cons, the_eq_cons] MRS msubst] 1
THEN REPEAT (assume_tac 1));
@@ -285,19 +285,19 @@
"!!k. [| k:nat; m:nat |] ==> \
\ ALL A B. A eqpoll k #+ m & k lepoll B & B<=A --> A-B lepoll m";
by (eres_inst_tac [("n","k")] nat_induct 1);
-by (simp_tac (AC_ss addsimps [add_0]) 1);
-by (fast_tac (AC_cs addIs [eqpoll_imp_lepoll RS
+by (simp_tac (!simpset addsimps [add_0]) 1);
+by (fast_tac (!claset addIs [eqpoll_imp_lepoll RS
(Diff_subset RS subset_imp_lepoll RS lepoll_trans)]) 1);
by (REPEAT (resolve_tac [allI,impI] 1));
by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (eres_inst_tac [("x","A - {xa}")] allE 1);
by (eres_inst_tac [("x","B - {xa}")] allE 1);
by (etac impE 1);
-by (asm_full_simp_tac (AC_ss addsimps [add_succ]) 1);
-by (fast_tac (AC_cs addSIs [Diff_sing_eqpoll, lepoll_Diff_sing]) 1);
+by (asm_full_simp_tac (!simpset addsimps [add_succ]) 1);
+by (fast_tac (!claset addSIs [Diff_sing_eqpoll, lepoll_Diff_sing]) 1);
by (res_inst_tac [("P","%z. z lepoll m")] subst 1 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val eqpoll_sum_imp_Diff_lepoll_lemma = result();
goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) lepoll B; \
@@ -306,7 +306,7 @@
by (dresolve_tac [add_succ RS ssubst] 1);
by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_lepoll_lemma] 1
THEN (REPEAT (assume_tac 1)));
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val eqpoll_sum_imp_Diff_lepoll = result();
(* ********************************************************************** *)
@@ -317,19 +317,19 @@
"!!k. [| k:nat; m:nat |] ==> \
\ ALL A B. A eqpoll k #+ m & k eqpoll B & B<=A --> A-B eqpoll m";
by (eres_inst_tac [("n","k")] nat_induct 1);
-by (fast_tac (AC_cs addSDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_0_is_0]
- addss (AC_ss addsimps [add_0])) 1);
+by (fast_tac (!claset addSDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_0_is_0]
+ addss (!simpset addsimps [add_0])) 1);
by (REPEAT (resolve_tac [allI,impI] 1));
by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1);
-by (fast_tac (AC_cs addSEs [eqpoll_imp_lepoll]) 1);
+by (fast_tac (!claset addSEs [eqpoll_imp_lepoll]) 1);
by (eres_inst_tac [("x","A - {xa}")] allE 1);
by (eres_inst_tac [("x","B - {xa}")] allE 1);
by (etac impE 1);
-by (fast_tac (AC_cs addSIs [Diff_sing_eqpoll,
+by (fast_tac (!claset addSIs [Diff_sing_eqpoll,
eqpoll_sym RSN (2, Diff_sing_eqpoll) RS eqpoll_sym]
- addss (AC_ss addsimps [add_succ])) 1);
+ addss (!simpset addsimps [add_succ])) 1);
by (res_inst_tac [("P","%z. z eqpoll m")] subst 1 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val eqpoll_sum_imp_Diff_eqpoll_lemma = result();
goal thy "!!k. [| A eqpoll succ(k #+ m); B<=A; succ(k) eqpoll B; \
@@ -338,7 +338,7 @@
by (dresolve_tac [add_succ RS ssubst] 1);
by (dresolve_tac [nat_succI RS eqpoll_sum_imp_Diff_eqpoll_lemma] 1
THEN (REPEAT (assume_tac 1)));
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val eqpoll_sum_imp_Diff_eqpoll = result();
(* ********************************************************************** *)
@@ -347,7 +347,7 @@
goal thy "!!w. [| x Int y = 0; w <= x Un y |] \
\ ==> w Int (x - {u}) = w - cons(u, w Int y)";
-by (fast_tac (AC_cs addSIs [equalityI] addEs [equals0D]) 1);
+by (fast_tac (!claset addSIs [equalityI] addEs [equals0D]) 1);
val w_Int_eq_w_Diff = result();
goal thy "!!w. [| w:{v:s_u(u, t_n, succ(l), y). a <= v}; \
@@ -357,8 +357,8 @@
\ |] ==> w Int (x - {u}) eqpoll m";
by (etac CollectE 1);
by (resolve_tac [w_Int_eq_w_Diff RS ssubst] 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSDs [s_u_subset RS subsetD]) 1);
-by (fast_tac (AC_cs addEs [equals0D] addSDs [bspec]
+by (fast_tac (!claset addSDs [s_u_subset RS subsetD]) 1);
+by (fast_tac (!claset addEs [equals0D] addSDs [bspec]
addDs [s_u_subset RS subsetD]
addSEs [eqpoll_sym RS cons_eqpoll_succ RS eqpoll_sym, in_s_u_imp_u_in]
addSIs [nat_succI, eqpoll_sum_imp_Diff_eqpoll]) 1);
@@ -372,7 +372,7 @@
goal thy
"!!z. [| z : xa Int (x - {u}); l eqpoll a; a <= y; x Int y = 0; u:x \
\ |] ==> cons(z, cons(u, a)) : {v: Pow(x Un y). v eqpoll succ(succ(l))}";
-by (fast_tac (AC_cs addSIs [cons_eqpoll_succ] addEs [equals0D, eqpoll_sym]) 1);
+by (fast_tac (!claset addSIs [cons_eqpoll_succ] addEs [equals0D, eqpoll_sym]) 1);
val cons_cons_in = result();
(* ********************************************************************** *)
@@ -396,17 +396,17 @@
by (rtac CollectI 1);
by (rtac lam_type 1);
by (rtac CollectI 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (rtac w_Int_eqpoll_m 1 THEN REPEAT (assume_tac 1));
-by (simp_tac AC_ss 1);
+by (Simp_tac 1);
by (REPEAT (resolve_tac [ballI, impI] 1));
by (eresolve_tac [w_Int_eqpoll_m RSN (2, eqpoll_m_not_empty) RS not_emptyE] 1
THEN REPEAT (assume_tac 1));
by (dresolve_tac [equalityD1 RS subsetD] 1 THEN (assume_tac 1));
by (dresolve_tac [cons_cons_in RSN (2, bspec)] 1 THEN REPEAT (assume_tac 1));
by (etac ex1_two_eq 1);
-by (fast_tac (AC_cs addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
-by (fast_tac (AC_cs addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
+by (fast_tac (!claset addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
+by (fast_tac (!claset addSEs [s_u_subset RS subsetD, in_s_u_imp_u_in]) 1);
val subset_s_u_lepoll_w = result();
goal thy "!!k. [| 0<k; k:nat |] ==> EX l:nat. k = succ(l)";
@@ -438,7 +438,7 @@
(* ********************************************************************** *)
goal thy "{x:Pow(X). x lepoll 0} = {0}";
-by (fast_tac (AC_cs addSDs [lepoll_0_is_0]
+by (fast_tac (!claset addSDs [lepoll_0_is_0]
addSIs [lepoll_refl, equalityI]) 1);
val subsets_lepoll_0_eq_unit = result();
@@ -447,12 +447,12 @@
by (resolve_tac [well_ord_infinite_subsets_eqpoll_X
RS (eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1
THEN (REPEAT (assume_tac 1)));
-by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage]) 1);
+by (fast_tac (!claset addSEs [bij_is_inj RS well_ord_rvimage]) 1);
val well_ord_subsets_eqpoll_n = result();
goal thy "!!n. n:nat ==> {z:Pow(y). z lepoll succ(n)} = \
\ {z:Pow(y). z lepoll n} Un {z:Pow(y). z eqpoll succ(n)}";
-by (fast_tac (ZF_cs addIs [le_refl, leI,
+by (fast_tac (!claset addIs [le_refl, leI,
le_imp_lepoll, equalityI]
addSDs [lepoll_succ_disj]
addSEs [nat_into_Ord, lepoll_trans, eqpoll_imp_lepoll]) 1);
@@ -460,7 +460,7 @@
goal thy "!!n. n:nat ==> \
\ {z:Pow(y). z lepoll n} Int {z:Pow(y). z eqpoll succ(n)} = 0";
-by (fast_tac (ZF_cs addSEs [eqpoll_sym RS eqpoll_imp_lepoll
+by (fast_tac (!claset addSEs [eqpoll_sym RS eqpoll_imp_lepoll
RS lepoll_trans RS succ_lepoll_natE]
addSIs [equals0I]) 1);
val Int_empty = result();
@@ -471,15 +471,15 @@
goalw thy [irrefl_def, trans_on_def, part_ord_def, linear_def, tot_ord_def]
"tot_ord({a},0)";
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
val tot_ord_unit = result();
goalw thy [wf_on_def, wf_def] "wf[{a}](0)";
-by (fast_tac (ZF_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val wf_on_unit = result();
goalw thy [well_ord_def] "well_ord({a},0)";
-by (simp_tac (ZF_ss addsimps [tot_ord_unit, wf_on_unit]) 1);
+by (simp_tac (!simpset addsimps [tot_ord_unit, wf_on_unit]) 1);
val well_ord_unit = result();
(* ********************************************************************** *)
@@ -489,22 +489,22 @@
goal thy "!!y r. [| well_ord(y,r); ~Finite(y); n:nat |] ==> \
\ EX R. well_ord({z:Pow(y). z lepoll n}, R)";
by (etac nat_induct 1);
-by (fast_tac (ZF_cs addSIs [well_ord_unit]
- addss (ZF_ss addsimps [subsets_lepoll_0_eq_unit])) 1);
+by (fast_tac (!claset addSIs [well_ord_unit]
+ addss (!simpset addsimps [subsets_lepoll_0_eq_unit])) 1);
by (etac exE 1);
by (eresolve_tac [well_ord_subsets_eqpoll_n RS exE] 1
THEN REPEAT (assume_tac 1));
-by (asm_simp_tac (ZF_ss addsimps [subsets_lepoll_succ]) 1);
+by (asm_simp_tac (!simpset addsimps [subsets_lepoll_succ]) 1);
by (dtac well_ord_radd 1 THEN (assume_tac 1));
by (eresolve_tac [Int_empty RS disj_Un_eqpoll_sum RS
(eqpoll_def RS def_imp_iff RS iffD1) RS exE] 1);
-by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage]) 1);
+by (fast_tac (!claset addSEs [bij_is_inj RS well_ord_rvimage]) 1);
val well_ord_subsets_lepoll_n = result();
goalw thy [LL_def, MM_def]
"!!x. t_n <= {v:Pow(x Un y). v eqpoll n} \
\ ==> LL(t_n, k, y) <= {z:Pow(y). z lepoll n}";
-by (fast_tac (AC_cs addSEs [RepFunE]
+by (fast_tac (!claset addSEs [RepFunE]
addIs [subset_imp_lepoll RS (eqpoll_imp_lepoll
RSN (2, lepoll_trans))]) 1);
val LL_subset = result();
@@ -526,16 +526,16 @@
\ t_n <= {v:Pow(x Un y). v eqpoll n}; \
\ v:LL(t_n, k, y) \
\ |] ==> EX! w. w:MM(t_n, k, y) & v<=w";
-by (step_tac (AC_cs addSEs [RepFunE]) 1);
+by (step_tac (!claset addSEs [RepFunE]) 1);
by (resolve_tac [lepoll_imp_eqpoll_subset RS exE] 1 THEN (assume_tac 1));
by (eres_inst_tac [("x","xa")] ballE 1);
-by (fast_tac (AC_cs addSEs [eqpoll_sym]) 2);
+by (fast_tac (!claset addSEs [eqpoll_sym]) 2);
by (res_inst_tac [("a","v")] ex1I 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (etac ex1E 1);
by (res_inst_tac [("x","v")] allE 1 THEN (assume_tac 1));
by (eres_inst_tac [("x","xb")] allE 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val unique_superset_in_MM = result();
(* ********************************************************************** *)
@@ -556,15 +556,15 @@
by (eresolve_tac [exists_proper_in_s_u RS bexE] 1
THEN REPEAT (assume_tac 1));
by (rewrite_goals_tac [MM_def, s_u_def]);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val exists_in_MM = result();
goalw thy [LL_def] "!!w. w : MM(t_n, k, y) ==> w Int y : LL(t_n, k, y)";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val Int_in_LL = result();
goalw thy [MM_def] "MM(t_n, k, y) <= t_n";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val MM_subset = result();
goal thy
@@ -579,10 +579,10 @@
by (res_inst_tac [("x","w Int y")] bexI 1);
by (etac Int_in_LL 2);
by (rewtac GG_def);
-by (asm_full_simp_tac (AC_ss addsimps [Int_in_LL]) 1);
+by (asm_full_simp_tac (!simpset addsimps [Int_in_LL]) 1);
by (eresolve_tac [unique_superset_in_MM RS the_equality2 RS ssubst] 1
THEN (assume_tac 1));
-by (REPEAT (fast_tac (AC_cs addEs [equals0D] addSEs [Int_in_LL]) 1));
+by (REPEAT (fast_tac (!claset addEs [equals0D] addSEs [Int_in_LL]) 1));
val exists_in_LL = result();
goalw thy [LL_def]
@@ -590,7 +590,7 @@
\ t_n <= {v:Pow(x Un y). v eqpoll n}; \
\ v : LL(t_n, k, y) |] \
\ ==> v = (THE x. x : MM(t_n, k, y) & v <= x) Int y";
-by (fast_tac (AC_cs addSEs [Int_in_LL,
+by (fast_tac (!claset addSEs [Int_in_LL,
unique_superset_in_MM RS the_equality2 RS ssubst]) 1);
val in_LL_eq_Int = result();
@@ -599,7 +599,7 @@
\ t_n <= {v:Pow(x Un y). v eqpoll n}; \
\ v : LL(t_n, k, y) |] \
\ ==> (THE x. x : MM(t_n, k, y) & v <= x) <= x Un y";
-by (fast_tac (AC_cs addSDs [unique_superset_in_MM RS theI RS conjunct1 RS
+by (fast_tac (!claset addSDs [unique_superset_in_MM RS theI RS conjunct1 RS
(MM_subset RS subsetD)]) 1);
val the_in_MM_subset = result();
@@ -608,13 +608,13 @@
\ t_n <= {v:Pow(x Un y). v eqpoll n}; \
\ v : LL(t_n, k, y) |] \
\ ==> GG(t_n, k, y) ` v <= x";
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
by (forward_tac [the_in_MM_subset] 1 THEN REPEAT (assume_tac 1));
by (dtac in_LL_eq_Int 1 THEN REPEAT (assume_tac 1));
by (rtac subsetI 1);
by (etac DiffE 1);
by (etac swap 1);
-by (fast_tac (AC_cs addEs [ssubst]) 1);
+by (fast_tac (!claset addEs [ssubst]) 1);
val GG_subset = result();
goal thy
@@ -649,12 +649,12 @@
goalw thy [MM_def]
"!!w. [| w : MM(t_n, k, y); t_n <= {v:Pow(x Un y). v eqpoll n} \
\ |] ==> w eqpoll n";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val in_MM_eqpoll_n = result();
goalw thy [LL_def, MM_def]
"!!w. w : LL(t_n, k, y) ==> k lepoll w";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val in_LL_eqpoll_n = result();
goalw thy [GG_def]
@@ -666,13 +666,14 @@
\ (GG(t_n, succ(k), y)) ` \
\ (converse(ordermap(LL(t_n, succ(k), y), S)) ` b) lepoll m";
by (rtac oallI 1);
-by (asm_full_simp_tac (AC_ss addsimps [ltD,
+by (asm_full_simp_tac (!simpset addsimps [ltD,
ordermap_bij RS bij_converse_bij RS bij_is_fun RS apply_type]) 1);
by (rtac eqpoll_sum_imp_Diff_lepoll 1);
-by (REPEAT (fast_tac (FOL_cs addSDs [ltD]
+by (REPEAT (fast_tac
+ (FOL_cs addSDs [ltD]
addSIs [eqpoll_sum_imp_Diff_lepoll, in_LL_eqpoll_n]
addEs [unique_superset_in_MM RS theI RS conjunct1 RS in_MM_eqpoll_n,
- in_LL_eq_Int RS equalityD1 RS (Int_lower1 RSN (2, subset_trans)),
+ in_LL_eq_Int RS equalityD1 RS (Int_lower1 RSN (2, subset_trans)),
ordermap_bij RS bij_converse_bij RS bij_is_fun RS apply_type]) 1));
val all_in_lepoll_m = result();
@@ -691,12 +692,12 @@
by (forward_tac [infinite_Un] 1 THEN (mp_tac 1));
by (REPEAT (eresolve_tac [exE, conjE] 1));
by (resolve_tac [well_ord_LL RS exE] 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSIs [nat_succI, add_type]) 1);
+by (fast_tac (!claset addSIs [nat_succI, add_type]) 1);
by (res_inst_tac [("x","ordertype(LL(T, succ(k), y), x)")] exI 1);
by (res_inst_tac [("x","lam b:ordertype(LL(T, succ(k), y), x). \
\ (GG(T, succ(k), y)) ` \
\ (converse(ordermap(LL(T, succ(k), y), x)) ` b)")] exI 1);
-by (simp_tac AC_ss 1);
+by (Simp_tac 1);
by (fast_tac (empty_cs addSIs [conjI, lam_funtype RS domain_of_fun]
addSEs [Ord_ordertype, all_in_lepoll_m, OUN_eq_x]) 1);
qed "AC16_WO4";
--- a/src/ZF/AC/AC16_lemmas.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC16_lemmas.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,28 +8,28 @@
open AC16_lemmas;
goal thy "!!a. a~:A ==> cons(a,A)-{a}=A";
-by (fast_tac (ZF_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val cons_Diff_eq = result();
goalw thy [lepoll_def] "1 lepoll X <-> (EX x. x:X)";
by (rtac iffI 1);
-by (fast_tac (ZF_cs addIs [inj_is_fun RS apply_type]) 1);
+by (fast_tac (!claset addIs [inj_is_fun RS apply_type]) 1);
by (etac exE 1);
by (res_inst_tac [("x","lam a:1. x")] exI 1);
-by (fast_tac (ZF_cs addSIs [lam_injective]) 1);
+by (fast_tac (!claset addSIs [lam_injective]) 1);
val nat_1_lepoll_iff = result();
goal thy "X eqpoll 1 <-> (EX x. X={x})";
by (rtac iffI 1);
by (etac eqpollE 1);
by (dresolve_tac [nat_1_lepoll_iff RS iffD1] 1);
-by (fast_tac (ZF_cs addSIs [lepoll_1_is_sing]) 1);
-by (fast_tac (ZF_cs addSIs [singleton_eqpoll_1]) 1);
+by (fast_tac (!claset addSIs [lepoll_1_is_sing]) 1);
+by (fast_tac (!claset addSIs [singleton_eqpoll_1]) 1);
val eqpoll_1_iff_singleton = result();
goalw thy [succ_def]
"!!x. [| x eqpoll n; y~:x |] ==> cons(y,x) eqpoll succ(n)";
-by (fast_tac (ZF_cs addSEs [cons_eqpoll_cong, mem_irrefl]) 1);
+by (fast_tac (!claset addSEs [cons_eqpoll_cong, mem_irrefl]) 1);
val cons_eqpoll_succ = result();
goal thy "{Y:Pow(X). Y eqpoll 1} = {{x}. x:X}";
@@ -37,23 +37,23 @@
by (rtac subsetI 1);
by (etac CollectE 1);
by (dresolve_tac [eqpoll_1_iff_singleton RS iffD1] 1);
-by (fast_tac (AC_cs addSIs [RepFunI]) 1);
+by (fast_tac (!claset addSIs [RepFunI]) 1);
by (rtac subsetI 1);
by (etac RepFunE 1);
by (rtac CollectI 1);
-by (fast_tac AC_cs 1);
-by (fast_tac (AC_cs addSIs [singleton_eqpoll_1]) 1);
+by (Fast_tac 1);
+by (fast_tac (!claset addSIs [singleton_eqpoll_1]) 1);
val subsets_eqpoll_1_eq = result();
goalw thy [eqpoll_def, bij_def] "X eqpoll {{x}. x:X}";
by (res_inst_tac [("x","lam x:X. {x}")] exI 1);
by (rtac IntI 1);
by (rewrite_goals_tac [inj_def, surj_def]);
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSIs [lam_type, RepFunI]
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSIs [lam_type, RepFunI]
addIs [singleton_eq_iff RS iffD1]) 1);
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSIs [lam_type]) 1);
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSIs [lam_type]) 1);
val eqpoll_RepFun_sing = result();
goal thy "{Y:Pow(X). Y eqpoll 1} eqpoll X";
@@ -65,7 +65,7 @@
\ ==> (LEAST i. i:y) : y";
by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS
succ_lepoll_imp_not_empty RS not_emptyE] 1);
-by (fast_tac (AC_cs addIs [LeastI]
+by (fast_tac (!claset addIs [LeastI]
addSDs [InfCard_is_Card RS Card_is_Ord, PowD RS subsetD]
addEs [Ord_in_Ord]) 1);
val InfCard_Least_in = result();
@@ -78,34 +78,34 @@
by (res_inst_tac [("d","%z. cons(fst(z),snd(z))")] lam_injective 1);
by (rtac SigmaI 1);
by (etac CollectE 1);
-by (asm_full_simp_tac AC_ss 3);
+by (Asm_full_simp_tac 3);
by (rtac equalityI 3);
-by (fast_tac AC_cs 4);
+by (Fast_tac 4);
by (rtac subsetI 3);
by (etac consE 3);
-by (fast_tac AC_cs 4);
+by (Fast_tac 4);
by (rtac CollectI 2);
-by (fast_tac AC_cs 2);
+by (Fast_tac 2);
by (resolve_tac [PowD RS subsetD] 1 THEN (assume_tac 1));
-by (REPEAT (fast_tac (AC_cs addSIs [Diff_sing_eqpoll]
+by (REPEAT (fast_tac (!claset addSIs [Diff_sing_eqpoll]
addIs [InfCard_Least_in]) 1));
val subsets_lepoll_lemma1 = result();
val prems = goal thy "(!!y. y:z ==> Ord(y)) ==> z <= succ(Union(z))";
by (rtac subsetI 1);
by (res_inst_tac [("Q","ALL y:z. y<=x")] (excluded_middle RS disjE) 1);
-by (fast_tac (AC_cs addSIs [equalityI]) 2);
+by (fast_tac (!claset addSIs [equalityI]) 2);
by (etac swap 1);
by (rtac ballI 1);
by (rtac Ord_linear_le 1);
by (dtac le_imp_subset 3 THEN (assume_tac 3));
-by (fast_tac (AC_cs addDs prems) 1);
-by (fast_tac (AC_cs addDs prems) 1);
-by (fast_tac (AC_cs addSEs [leE,ltE]) 1);
+by (fast_tac (!claset addDs prems) 1);
+by (fast_tac (!claset addDs prems) 1);
+by (fast_tac (!claset addSEs [leE,ltE]) 1);
val set_of_Ord_succ_Union = result();
goal thy "!!i. j<=i ==> i ~: j";
-by (fast_tac (ZF_cs addSEs [mem_irrefl]) 1);
+by (fast_tac (!claset addSEs [mem_irrefl]) 1);
val subset_not_mem = result();
val prems = goal thy "(!!y. y:z ==> Ord(y)) ==> succ(Union(z)) ~: z";
@@ -114,31 +114,31 @@
val succ_Union_not_mem = result();
goal thy "Union(cons(succ(Union(z)),z)) = succ(Union(z))";
-by (fast_tac (AC_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
val Union_cons_eq_succ_Union = result();
goal thy "!!i. [| Ord(i); Ord(j) |] ==> i Un j = i | i Un j = j";
-by (fast_tac (AC_cs addSDs [le_imp_subset] addIs [equalityI]
+by (fast_tac (!claset addSDs [le_imp_subset] addIs [equalityI]
addEs [Ord_linear_le]) 1);
val Un_Ord_disj = result();
goal thy "!!X. x:X ==> Union(X) = x Un Union(X-{x})";
-by (fast_tac (AC_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
val Union_eq_Un = result();
goal thy "!!n. n:nat ==> \
\ ALL z. (ALL y:z. Ord(y)) & z eqpoll n & z~=0 --> Union(z) : z";
by (etac nat_induct 1);
-by (fast_tac (AC_cs addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1);
+by (fast_tac (!claset addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1);
by (REPEAT (resolve_tac [allI, impI] 1));
by (etac natE 1);
-by (fast_tac (AC_cs addSDs [eqpoll_1_iff_singleton RS iffD1]
+by (fast_tac (!claset addSDs [eqpoll_1_iff_singleton RS iffD1]
addSIs [Union_singleton]) 1);
by (hyp_subst_tac 1);
by (REPEAT (eresolve_tac [conjE, not_emptyE] 1));
by (eres_inst_tac [("x","z-{xb}")] allE 1);
by (etac impE 1);
-by (fast_tac (AC_cs addSEs [Diff_sing_eqpoll,
+by (fast_tac (!claset addSEs [Diff_sing_eqpoll,
Diff_sing_eqpoll RS eqpoll_succ_imp_not_empty]) 1);
by (resolve_tac [Union_eq_Un RSN (2, subst_elem)] 1 THEN (assume_tac 2));
by (forward_tac [bspec] 1 THEN (assume_tac 1));
@@ -161,14 +161,14 @@
by (resolve_tac [Limit_has_succ RS ltE] 1 THEN (assume_tac 3));
by (etac InfCard_is_Limit 1);
by (excluded_middle_tac "z=0" 1);
-by (fast_tac (AC_cs addSIs [InfCard_is_Limit RS Limit_has_0]
- addIs [Union_0 RS ssubst]) 2);
+by (fast_tac (!claset addSIs [InfCard_is_Limit RS Limit_has_0]
+ addss (!simpset)) 2);
by (resolve_tac
[PowD RS subsetD RS (InfCard_is_Card RS Card_is_Ord RSN (2, ltI))] 1
THEN (TRYALL assume_tac));
-by (fast_tac (AC_cs addSIs [Union_in]
- addSEs [PowD RS subsetD RSN (2,
- InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1);
+by (fast_tac (!claset addSIs [Union_in]
+ addSEs [PowD RS subsetD RSN
+ (2, InfCard_is_Card RS Card_is_Ord RS Ord_in_Ord)]) 1);
val succ_Union_in_x = result();
goalw thy [lepoll_def] "!!X. [| InfCard(x); n:nat |] ==> \
@@ -179,14 +179,14 @@
by (res_inst_tac [("d","%z. z-{Union(z)}")] lam_injective 1);
by (resolve_tac [Union_cons_eq_succ_Union RS ssubst] 2);
by (rtac cons_Diff_eq 2);
-by (fast_tac (AC_cs addSDs [InfCard_is_Card RS Card_is_Ord]
+by (fast_tac (!claset addSDs [InfCard_is_Card RS Card_is_Ord]
addEs [Ord_in_Ord] addSIs [succ_Union_not_mem]) 2);
by (rtac CollectI 1);
-by (fast_tac (AC_cs addSEs [cons_eqpoll_succ]
+by (fast_tac (!claset addSEs [cons_eqpoll_succ]
addSIs [succ_Union_not_mem]
addSDs [InfCard_is_Card RS Card_is_Ord]
addEs [Ord_in_Ord]) 2);
-by (fast_tac (AC_cs addSIs [succ_Union_in_x, nat_succI]) 1);
+by (fast_tac (!claset addSIs [succ_Union_in_x, nat_succI]) 1);
val succ_lepoll_succ_succ = result();
goal thy "!!X. [| InfCard(X); n:nat |] \
@@ -202,18 +202,18 @@
by (resolve_tac [InfCard_is_Card RS Card_cardinal_eq RS ssubst] 2
THEN (REPEAT (assume_tac 2)));
by (eresolve_tac [eqpoll_refl RS prod_eqpoll_cong RS eqpoll_imp_lepoll] 1);
-by (fast_tac (AC_cs addEs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans]
+by (fast_tac (!claset addEs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans]
addSIs [succ_lepoll_succ_succ]) 1);
val subsets_eqpoll_X = result();
goalw thy [surj_def] "!!f. [| f:surj(A,B); y<=B |] \
\ ==> f``(converse(f)``y) = y";
-by (fast_tac (AC_cs addSIs [equalityI] addDs [apply_equality2]
+by (fast_tac (!claset addSIs [equalityI] addDs [apply_equality2]
addEs [apply_iff RS iffD2]) 1);
val image_vimage_eq = result();
goal thy "!!f. [| f:inj(A,B); y<=A |] ==> converse(f)``(f``y) = y";
-by (fast_tac (AC_cs addSIs [equalityI] addSEs [inj_is_fun RS apply_Pair]
+by (fast_tac (!claset addSIs [equalityI] addSEs [inj_is_fun RS apply_Pair]
addDs [inj_equality]) 1);
val vimage_image_eq = result();
@@ -222,20 +222,20 @@
by (etac exE 1);
by (res_inst_tac [("x","lam X:{Y:Pow(A). EX f. f : bij(Y, n)}. f``X")] exI 1);
by (res_inst_tac [("d","%Z. converse(f)``Z")] lam_bijective 1);
-by (fast_tac (AC_cs
+by (fast_tac (!claset
addSIs [bij_is_inj RS restrict_bij RS bij_converse_bij RS comp_bij]
addSEs [bij_is_fun RS fun_is_rel RS image_subset RS PowI]) 1);
-by (fast_tac (AC_cs addSIs [bij_converse_bij RS bij_is_inj RS restrict_bij
+by (fast_tac (!claset addSIs [bij_converse_bij RS bij_is_inj RS restrict_bij
RS bij_converse_bij RS comp_bij]
addSEs [bij_converse_bij RS bij_is_fun RS fun_is_rel
RS image_subset RS PowI]) 1);
-by (fast_tac (AC_cs addSEs [bij_is_inj RS vimage_image_eq]) 1);
-by (fast_tac (AC_cs addSEs [bij_is_surj RS image_vimage_eq]) 1);
+by (fast_tac (!claset addSEs [bij_is_inj RS vimage_image_eq]) 1);
+by (fast_tac (!claset addSEs [bij_is_surj RS image_vimage_eq]) 1);
val subsets_eqpoll = result();
goalw thy [WO2_def] "!!X. WO2 ==> EX a. Card(a) & X eqpoll a";
by (REPEAT (eresolve_tac [allE,exE,conjE] 1));
-by (fast_tac (AC_cs addSEs [well_ord_Memrel RS well_ord_cardinal_eqpoll RS
+by (fast_tac (!claset addSEs [well_ord_Memrel RS well_ord_cardinal_eqpoll RS
(eqpoll_sym RSN (2, eqpoll_trans)) RS eqpoll_sym]
addSIs [Card_cardinal]) 1);
val WO2_imp_ex_Card = result();
@@ -245,7 +245,7 @@
val lepoll_infinite = result();
goalw thy [InfCard_def] "!!X. [| ~Finite(X); Card(X) |] ==> InfCard(X)";
-by (fast_tac (AC_cs addSEs [Card_is_Ord RS nat_le_infinite_Ord]) 1);
+by (fast_tac (!claset addSEs [Card_is_Ord RS nat_le_infinite_Ord]) 1);
val infinite_Card_is_InfCard = result();
goal thy "!!X n. [| WO2; n:nat; ~Finite(X) |] \
@@ -261,7 +261,7 @@
val WO2_infinite_subsets_eqpoll_X = result();
goal thy "!!X. well_ord(X,R) ==> EX a. Card(a) & X eqpoll a";
-by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll RS eqpoll_sym]
+by (fast_tac (!claset addSEs [well_ord_cardinal_eqpoll RS eqpoll_sym]
addSIs [Card_cardinal]) 1);
val well_ord_imp_ex_Card = result();
--- a/src/ZF/AC/AC16_lemmas.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC16_lemmas.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,3 @@
(*Dummy theory to document dependencies *)
-AC16_lemmas = AC_Equiv + Hartog + first
+AC16_lemmas = AC_Equiv + Hartog
--- a/src/ZF/AC/AC17_AC1.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC17_AC1.ML Fri Jan 03 15:01:55 1997 +0100
@@ -34,7 +34,7 @@
by (rtac ballI 1);
by (etac swap 1);
by (rtac impI 1);
-by (fast_tac (AC_cs addSIs [restrict_type]) 1);
+by (fast_tac (!claset addSIs [restrict_type]) 1);
val not_AC1_imp_ex = result();
goal thy "!!x. [| ALL f:Pow(x) - {0} -> x. EX u: Pow(x) - {0}. f`u~:u; \
@@ -50,25 +50,25 @@
by (etac notE 1);
by (rtac Pi_type 1 THEN (assume_tac 1));
by (resolve_tac [apply_type RSN (2, subsetD)] 1 THEN TRYALL assume_tac);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val lemma1 = result();
goal thy "!!x. ~ (EX f: Pow(x)-{0}->x. x - F(f) = 0) \
\ ==> (lam f: Pow(x)-{0}->x. x - F(f)) \
\ : (Pow(x) -{0} -> x) -> Pow(x) - {0}";
-by (fast_tac (AC_cs addSIs [lam_type] addIs [equalityI]
+by (fast_tac (!claset addSIs [lam_type] addIs [equalityI]
addSDs [Diff_eq_0_iff RS iffD1]) 1);
val lemma2 = result();
goal thy "!!f. [| f`Z : Z; Z:Pow(x)-{0} |] ==> \
\ (lam X:Pow(x)-{0}. {f`X})`Z : Pow(Z)-{0}";
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSDs [equals0D]) 1);
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSDs [equals0D]) 1);
val lemma3 = result();
goal thy "!!z. EX f:F. f`((lam f:F. Q(f))`f) : (lam f:F. Q(f))`f \
\ ==> EX f:F. f`Q(f) : Q(f)";
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val lemma4 = result();
goalw thy [AC17_def] "!!Z. AC17 ==> AC1";
@@ -88,6 +88,6 @@
by (dresolve_tac [beta RS sym RSN (2, subst_elem)] 1);
by (assume_tac 1);
by (dtac lemma3 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem),
+by (fast_tac (!claset addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem),
f_subset_imp_HH_subset] addSEs [mem_irrefl]) 1);
qed "AC17_AC1";
--- a/src/ZF/AC/AC18_AC19.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC18_AC19.ML Fri Jan 03 15:01:55 1997 +0100
@@ -25,19 +25,19 @@
by (rtac subsetI 1);
by (eres_inst_tac [("x","{{b:B(a). x:X(a,b)}. a:A}")] allE 1);
by (etac impE 1);
-by (fast_tac (AC_cs addSEs [RepFunE] addSDs [INT_E]
+by (fast_tac (!claset addSEs [RepFunE] addSDs [INT_E]
addEs [UN_E, sym RS equals0D]) 1);
by (etac exE 1);
by (rtac UN_I 1);
-by (fast_tac (AC_cs addSEs [PROD_subsets]) 1);
-by (simp_tac AC_ss 1);
-by (fast_tac (FOL_cs addSEs [not_emptyE] addDs [RepFunI RSN (2, apply_type)]
+by (fast_tac (!claset addSEs [PROD_subsets]) 1);
+by (Simp_tac 1);
+by (fast_tac (!claset addSEs [not_emptyE] addDs [RepFunI RSN (2, apply_type)]
addEs [CollectD2] addSIs [INT_I]) 1);
val lemma_AC18 = result();
val [prem] = goalw thy (AC18_def::AC_defs) "AC1 ==> AC18";
by (resolve_tac [prem RS revcut_rl] 1);
-by (fast_tac (AC_cs addSEs [lemma_AC18, UN_E, not_emptyE, apply_type]
+by (fast_tac (!claset addSEs [lemma_AC18, UN_E, not_emptyE, apply_type]
addSIs [equalityI, INT_I, UN_I]) 1);
qed "AC1_AC18";
@@ -48,7 +48,7 @@
val [prem] = goalw thy [AC18_def, AC19_def] "AC18 ==> AC19";
by (rtac allI 1);
by (res_inst_tac [("B1","%x.x")] (forall_elim_vars 0 prem RS revcut_rl) 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
qed "AC18_AC19";
(* ********************************************************************** *)
@@ -57,7 +57,7 @@
goalw thy [u_def]
"!!A. [| A ~= 0; 0 ~: A |] ==> {u_(a). a:A} ~= 0 & 0 ~: {u_(a). a:A}";
-by (fast_tac (AC_cs addSIs [not_emptyI, RepFunI]
+by (fast_tac (!claset addSIs [not_emptyI, RepFunI]
addSEs [not_emptyE, RepFunE]
addSDs [sym RS (RepFun_eq_0_iff RS iffD1)]) 1);
val RepRep_conj = result();
@@ -66,17 +66,17 @@
by (hyp_subst_tac 1);
by (rtac subst_elem 1 THEN (assume_tac 1));
by (rtac equalityI 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (rtac subsetI 1);
by (excluded_middle_tac "x=0" 1);
-by (fast_tac AC_cs 1);
-by (fast_tac (AC_cs addEs [notE, subst_elem] addSIs [equalityI]) 1);
+by (Fast_tac 1);
+by (fast_tac (!claset addEs [notE, subst_elem] addSIs [equalityI]) 1);
val lemma1_1 = result();
goalw thy [u_def]
"!!a. [| f`(u_(a)) ~: a; f: (PROD B:{u_(a). a:A}. B); a:A |] \
\ ==> f`(u_(a))-{0} : a";
-by (fast_tac (AC_cs addSEs [RepFunI, RepFunE, lemma1_1]
+by (fast_tac (!claset addSEs [RepFunI, RepFunE, lemma1_1]
addSDs [apply_type]) 1);
val lemma1_2 = result();
@@ -87,35 +87,35 @@
by (rtac lam_type 1);
by (split_tac [expand_if] 1);
by (rtac conjI 1);
-by (fast_tac AC_cs 1);
-by (fast_tac (AC_cs addSEs [lemma1_2]) 1);
+by (Fast_tac 1);
+by (fast_tac (!claset addSEs [lemma1_2]) 1);
val lemma1 = result();
goalw thy [u_def] "!!a. a~=0 ==> 0: (UN b:u_(a). b)";
-by (fast_tac (AC_cs addSEs [not_emptyE] addSIs [UN_I, RepFunI]) 1);
+by (fast_tac (!claset addSEs [not_emptyE] addSIs [UN_I, RepFunI]) 1);
val lemma2_1 = result();
goal thy "!!A C. [| A~=0; 0~:A |] ==> (INT x:{u_(a). a:A}. UN b:x. b) ~= 0";
by (etac not_emptyE 1);
by (res_inst_tac [("a","0")] not_emptyI 1);
-by (fast_tac (AC_cs addSIs [INT_I, RepFunI, lemma2_1] addSEs [RepFunE]) 1);
+by (fast_tac (!claset addSIs [INT_I, RepFunI, lemma2_1] addSEs [RepFunE]) 1);
val lemma2 = result();
goal thy "!!F. (UN f:F. P(f)) ~= 0 ==> F ~= 0";
-by (fast_tac (AC_cs addSEs [not_emptyE]) 1);
+by (fast_tac (!claset addSEs [not_emptyE]) 1);
val lemma3 = result();
goalw thy AC_defs "!!Z. AC19 ==> AC1";
by (REPEAT (resolve_tac [allI,impI] 1));
by (excluded_middle_tac "A=0" 1);
-by (fast_tac (AC_cs addSIs [empty_fun]) 2);
+by (fast_tac (!claset addSIs [empty_fun]) 2);
by (eres_inst_tac [("x","{u_(a). a:A}")] allE 1);
by (etac impE 1);
by (etac RepRep_conj 1 THEN (assume_tac 1));
by (rtac lemma1 1);
by (dtac lemma2 1 THEN (assume_tac 1));
by (dres_inst_tac [("P","%x. x~=0")] subst 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [lemma3 RS not_emptyE]) 1);
+by (fast_tac (!claset addSEs [lemma3 RS not_emptyE]) 1);
qed "AC19_AC1";
--- a/src/ZF/AC/AC1_AC17.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC1_AC17.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,7 +8,7 @@
goal thy "!!f. f : (PROD X:Pow(A) - {0}. X) ==> f : (Pow(A) - {0} -> A)";
by (rtac Pi_type 1 THEN (assume_tac 1));
by (dtac apply_type 1 THEN (assume_tac 1));
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val lemma1 = result();
goalw thy AC_defs "!!Z. AC1 ==> AC17";
@@ -16,10 +16,10 @@
by (rtac ballI 1);
by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
by (etac impE 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (etac exE 1);
by (rtac bexI 1);
by (etac lemma1 2);
by (rtac apply_type 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSDs [lemma1] addSEs [apply_type]) 1);
+by (fast_tac (!claset addSDs [lemma1] addSEs [apply_type]) 1);
qed "AC1_AC17";
--- a/src/ZF/AC/AC1_WO2.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC1_WO2.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,12 +13,12 @@
by (resolve_tac [bij_Least_HH_x RS bij_converse_bij] 1);
by (rtac f_subsets_imp_UN_HH_eq_x 1);
by (resolve_tac [lam_type RS apply_type] 1 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSDs [equals0D, prem RS apply_type]) 1);
-by (fast_tac (AC_cs addSIs [prem RS Pi_weaken_type]) 1);
+by (fast_tac (!claset addSDs [equals0D, prem RS apply_type]) 1);
+by (fast_tac (!claset addSIs [prem RS Pi_weaken_type]) 1);
val lemma1 = uresult() |> standard;
goalw thy [AC1_def, WO2_def, eqpoll_def] "!!Z. AC1 ==> WO2";
by (rtac allI 1);
by (eres_inst_tac [("x","Pow(A)-{0}")] allE 1);
-by (fast_tac (AC_cs addSDs [lemma1] addSIs [Ord_Least]) 1);
+by (fast_tac (!claset addSDs [lemma1] addSIs [Ord_Least]) 1);
qed "AC1_WO2";
--- a/src/ZF/AC/AC2_AC6.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC2_AC6.ML Fri Jan 03 15:01:55 1997 +0100
@@ -16,12 +16,12 @@
goal thy "!!B. [| B:A; f:(PROD X:A. X); 0~:A |] \
\ ==> {f`B} <= B Int {f`C. C:A}";
-by (fast_tac (ZF_cs addSEs [apply_type]) 1);
+by (fast_tac (!claset addSEs [apply_type]) 1);
val lemma1 = result();
goalw thy [pairwise_disjoint_def]
"!!A. [| pairwise_disjoint(A); B:A; C:A; D:B; D:C |] ==> f`B = f`C";
-by (fast_tac (ZF_cs addSEs [equals0D]) 1);
+by (fast_tac (!claset addSEs [equals0D]) 1);
val lemma2 = result();
goalw thy AC_defs "!!Z. AC1 ==> AC2";
@@ -30,7 +30,7 @@
by (REPEAT (eresolve_tac [asm_rl,conjE,allE,exE,impE] 1));
by (REPEAT (resolve_tac [exI,ballI,equalityI] 1));
by (rtac lemma1 2 THEN (REPEAT (assume_tac 2)));
-by (fast_tac (AC_cs addSEs [RepFunE, lemma2] addEs [apply_type]) 1);
+by (fast_tac (!claset addSEs [RepFunE, lemma2] addEs [apply_type]) 1);
qed "AC1_AC2";
@@ -39,22 +39,22 @@
(* ********************************************************************** *)
goal thy "!!A. 0~:A ==> 0 ~: {B*{B}. B:A}";
-by (fast_tac (AC_cs addSDs [sym RS (Sigma_empty_iff RS iffD1)]
+by (fast_tac (!claset addSDs [sym RS (Sigma_empty_iff RS iffD1)]
addSEs [RepFunE, equals0D]) 1);
val lemma1 = result();
goal thy "!!A. [| X*{X} Int C = {y}; X:A |] \
\ ==> (THE y. X*{X} Int C = {y}): X*A";
by (rtac subst_elem 1);
-by (fast_tac (ZF_cs addSIs [the_equality]
+by (fast_tac (!claset addSIs [the_equality]
addSEs [sym RS trans RS (singleton_eq_iff RS iffD1)]) 2);
-by (fast_tac (AC_cs addSEs [equalityE, make_elim singleton_subsetD]) 1);
+by (fast_tac (!claset addSEs [equalityE, make_elim singleton_subsetD]) 1);
val lemma2 = result();
goal thy "!!A. ALL D:{E*{E}. E:A}. EX y. D Int C = {y} \
\ ==> (lam x:A. fst(THE z. (x*{x} Int C = {z}))) : \
\ (PROD X:A. X) ";
-by (fast_tac (FOL_cs addSEs [lemma2]
+by (fast_tac (!claset addSEs [lemma2]
addSIs [lam_type, RepFunI, fst_type]
addSDs [bspec]) 1);
val lemma3 = result();
@@ -62,8 +62,8 @@
goalw thy (AC_defs@AC_aux_defs) "!!Z. AC2 ==> AC1";
by (REPEAT (resolve_tac [allI, impI] 1));
by (REPEAT (eresolve_tac [allE, impE] 1));
-by (fast_tac (AC_cs addSEs [lemma3]) 2);
-by (fast_tac (AC_cs addSIs [lemma1, equals0I]) 1);
+by (fast_tac (!claset addSEs [lemma3]) 2);
+by (fast_tac (!claset addSIs [lemma1, equals0I]) 1);
qed "AC2_AC1";
@@ -72,13 +72,13 @@
(* ********************************************************************** *)
goal thy "!!R. 0 ~: {R``{x}. x:domain(R)}";
-by (fast_tac (AC_cs addEs [sym RS equals0D]) 1);
+by (fast_tac (!claset addEs [sym RS equals0D]) 1);
val lemma = result();
goalw thy AC_defs "!!Z. AC1 ==> AC4";
by (REPEAT (resolve_tac [allI, impI] 1));
by (REPEAT (eresolve_tac [allE, lemma RSN (2, impE), exE] 1));
-by (fast_tac (AC_cs addSIs [lam_type] addSEs [apply_type]) 1);
+by (fast_tac (!claset addSIs [lam_type] addSEs [apply_type]) 1);
qed "AC1_AC4";
@@ -87,25 +87,25 @@
(* ********************************************************************** *)
goal thy "!!f. f:A->B ==> (UN z:A. {z}*f`z) <= A*Union(B)";
-by (fast_tac (ZF_cs addSDs [apply_type]) 1);
+by (fast_tac (!claset addSDs [apply_type]) 1);
val lemma1 = result();
goal thy "!!f. domain(UN z:A. {z}*f(z)) = {a:A. f(a)~=0}";
-by (fast_tac (ZF_cs addIs [equalityI]
+by (fast_tac (!claset addIs [equalityI]
addSEs [not_emptyE]
addSIs [not_emptyI]
addDs [range_type]) 1);
val lemma2 = result();
goal thy "!!f. x:A ==> (UN z:A. {z}*f(z))``{x} = f(x)";
-by (fast_tac (ZF_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
val lemma3 = result();
goalw thy AC_defs "!!Z. AC4 ==> AC3";
by (REPEAT (resolve_tac [allI,ballI] 1));
by (REPEAT (eresolve_tac [allE,impE] 1));
by (etac lemma1 1);
-by (asm_full_simp_tac (AC_ss addsimps [lemma2, lemma3]
+by (asm_full_simp_tac (!simpset addsimps [lemma2, lemma3]
addcongs [Pi_cong]) 1);
qed "AC4_AC3";
@@ -114,16 +114,16 @@
(* ********************************************************************** *)
goal thy "!!A. b~:A ==> (PROD x:{a:A. id(A)`a~=b}. id(A)`x) = (PROD x:A. x)";
-by (asm_full_simp_tac (AC_ss addsimps [id_def] addcongs [Pi_cong]) 1);
+by (asm_full_simp_tac (!simpset addsimps [id_def] addcongs [Pi_cong]) 1);
by (res_inst_tac [("b","A")] subst_context 1);
-by (fast_tac (AC_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val lemma = result();
goalw thy AC_defs "!!Z. AC3 ==> AC1";
by (REPEAT (resolve_tac [allI, impI] 1));
by (REPEAT (eresolve_tac [allE, ballE] 1));
-by (fast_tac (AC_cs addSIs [id_type]) 2);
-by (fast_tac (AC_cs addEs [lemma RS subst]) 1);
+by (fast_tac (!claset addSIs [id_type]) 2);
+by (fast_tac (!claset addEs [lemma RS subst]) 1);
qed "AC3_AC1";
(* ********************************************************************** *)
@@ -137,13 +137,13 @@
by (etac exE 1);
by (rtac bexI 1);
by (rtac Pi_type 2 THEN (assume_tac 2));
-by (fast_tac (ZF_cs addSDs [apply_type]
+by (fast_tac (!claset addSDs [apply_type]
addSEs [fun_is_rel RS converse_type RS subsetD RS SigmaD2]) 2);
by (rtac ballI 1);
by (rtac apply_equality 1 THEN (assume_tac 2));
by (etac domainE 1);
by (forward_tac [range_type] 1 THEN (assume_tac 1));
-by (fast_tac (ZF_cs addDs [apply_equality]) 1);
+by (fast_tac (!claset addDs [apply_equality]) 1);
qed "AC4_AC5";
@@ -152,24 +152,24 @@
(* ********************************************************************** *)
goal thy "!!A. R <= A*B ==> (lam x:R. fst(x)) : R -> A";
-by (fast_tac (ZF_cs addSIs [lam_type, fst_type]) 1);
+by (fast_tac (!claset addSIs [lam_type, fst_type]) 1);
val lemma1 = result();
goalw thy [range_def] "!!A. R <= A*B ==> range(lam x:R. fst(x)) = domain(R)";
by (rtac equalityI 1);
-by (fast_tac (AC_cs addSEs [lamE]
+by (fast_tac (!claset addSEs [lamE]
addEs [subst_elem]
addSDs [Pair_fst_snd_eq]) 1);
by (rtac subsetI 1);
by (etac domainE 1);
by (rtac domainI 1);
-by (fast_tac (AC_cs addSEs [lamI RS subst_elem] addIs [fst_conv RS ssubst]) 1);
+by (fast_tac (!claset addSEs [lamI RS subst_elem] addIs [fst_conv RS ssubst]) 1);
val lemma2 = result();
goal thy "!!A. [| EX f: A->C. P(f,domain(f)); A=B |] ==> EX f: B->C. P(f,B)";
by (etac bexE 1);
by (forward_tac [domain_of_fun] 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val lemma3 = result();
goal thy "!!g. [| R <= A*B; g: C->R; ALL x:C. (lam z:R. fst(z))` (g`x) = x |] \
@@ -180,7 +180,7 @@
by (rtac imageI 1);
by (resolve_tac [subsetD RS Pair_fst_snd_eq RSN (2, subst_elem)] 1
THEN (REPEAT (assume_tac 1)));
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val lemma4 = result();
goalw thy AC_defs "!!Z. AC5 ==> AC4";
@@ -188,7 +188,7 @@
by (REPEAT (eresolve_tac [allE,ballE] 1));
by (eresolve_tac [lemma1 RSN (2, notE)] 2 THEN (assume_tac 2));
by (dresolve_tac [lemma2 RSN (2, lemma3)] 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [lemma4]) 1);
+by (fast_tac (!claset addSEs [lemma4]) 1);
qed "AC5_AC4";
@@ -197,6 +197,6 @@
(* ********************************************************************** *)
goalw thy AC_defs "AC1 <-> AC6";
-by (fast_tac (ZF_cs addDs [equals0D] addSEs [not_emptyE]) 1);
+by (fast_tac (!claset addDs [equals0D] addSEs [not_emptyE]) 1);
qed "AC1_iff_AC6";
--- a/src/ZF/AC/AC7_AC9.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC7_AC9.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,12 +13,12 @@
(* - Sigma_fun_space_eqpoll *)
(* ********************************************************************** *)
-goal ZF.thy "!!A. [| C~:A; B:A |] ==> B~=C";
-by (fast_tac ZF_cs 1);
+goal upair.thy "!!A. [| C~:A; B:A |] ==> B~=C";
+by (Fast_tac 1);
val mem_not_eq_not_mem = result();
goal thy "!!A. [| 0~:A; B:A |] ==> (nat->Union(A))*B ~= 0";
-by (fast_tac (ZF_cs addSDs [Sigma_empty_iff RS iffD1]
+by (fast_tac (!claset addSDs [Sigma_empty_iff RS iffD1]
addDs [fun_space_emptyD, mem_not_eq_not_mem]
addEs [equals0D]
addSIs [equals0I,UnionI]) 1);
@@ -33,7 +33,7 @@
goal thy "!!A. [| ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)); b:A \
\ |] ==> P(b)=R(b)";
by (dtac bspec 1 THEN (assume_tac 1));
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val if_eqE1 = result();
goal thy "!!A. ALL a:A. if(a=b, P(a), Q(a)) = if(a=b, R(a), S(a)) \
@@ -41,11 +41,11 @@
by (rtac ballI 1);
by (rtac impI 1);
by (dtac bspec 1 THEN (assume_tac 1));
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val if_eqE2 = result();
goal thy "!!A. [| (lam x:A. f(x))=(lam x:A. g(x)); a:A |] ==> f(a)=g(a)";
-by (fast_tac (ZF_cs addDs [subsetD]
+by (fast_tac (!claset addDs [subsetD]
addSIs [lamI]
addEs [equalityE, lamE]) 1);
val lam_eqE = result();
@@ -55,29 +55,27 @@
\ (lam n:nat. if(n=0, snd(g), fst(g)`(n #- 1)))) \
\ : inj((nat->Union(A))*C, (nat->Union(A)) ) ";
by (rtac CollectI 1);
-by (fast_tac (ZF_cs addSIs [lam_type,RepFunI,if_type,snd_type,apply_type,
+by (fast_tac (!claset addSIs [lam_type,RepFunI,if_type,snd_type,apply_type,
fst_type,diff_type,nat_succI,nat_0I]) 1);
by (REPEAT (resolve_tac [ballI, impI] 1));
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
by (REPEAT (etac SigmaE 1));
by (REPEAT (hyp_subst_tac 1));
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
by (rtac conjI 1);
by (dresolve_tac [nat_0I RSN (2, lam_eqE)] 2);
-by (asm_full_simp_tac AC_ss 2);
+by (Asm_full_simp_tac 2);
by (rtac fun_extension 1 THEN REPEAT (assume_tac 1));
by (dresolve_tac [nat_succI RSN (2, lam_eqE)] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (AC_ss addsimps [succ_not_0 RS if_not_P]) 1);
-by (fast_tac (AC_cs addSEs [diff_succ_succ RS (diff_0 RSN (2, trans)) RS subst]
- addSIs [nat_0I]) 1);
+by (asm_full_simp_tac (!simpset addsimps [succ_not_0 RS if_not_P]) 1);
val lemma = result();
goal thy "!!A. [| C:A; 0~:A |] ==> (nat->Union(A)) * C eqpoll (nat->Union(A))";
by (rtac eqpollI 1);
-by (fast_tac (ZF_cs addSEs [prod_lepoll_self, not_sym RS not_emptyE,
+by (fast_tac (!claset addSEs [prod_lepoll_self, not_sym RS not_emptyE,
subst_elem] addEs [swap]) 2);
by (rewtac lepoll_def);
-by (fast_tac (ZF_cs addSIs [lemma]) 1);
+by (fast_tac (!claset addSIs [lemma]) 1);
val Sigma_fun_space_eqpoll = result();
@@ -86,7 +84,7 @@
(* ********************************************************************** *)
goalw thy AC_defs "!!Z. AC6 ==> AC7";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "AC6_AC7";
(* ********************************************************************** *)
@@ -96,22 +94,22 @@
(* ********************************************************************** *)
goal thy "!!y. y: (PROD B:A. Y*B) ==> (lam B:A. snd(y`B)): (PROD B:A. B)";
-by (fast_tac (ZF_cs addSIs [lam_type, snd_type, apply_type]) 1);
+by (fast_tac (!claset addSIs [lam_type, snd_type, apply_type]) 1);
val lemma1_1 = result();
goal thy "!!A. y: (PROD B:{Y*C. C:A}. B) \
\ ==> (lam B:A. y`(Y*B)): (PROD B:A. Y*B)";
-by (fast_tac (ZF_cs addSIs [lam_type, apply_type]) 1);
+by (fast_tac (!claset addSIs [lam_type, apply_type]) 1);
val lemma1_2 = result();
goal thy "!!A. (PROD B:{(nat->Union(A))*C. C:A}. B) ~= 0 \
\ ==> (PROD B:A. B) ~= 0";
-by (fast_tac (ZF_cs addSIs [equals0I,lemma1_1, lemma1_2]
+by (fast_tac (!claset addSIs [equals0I,lemma1_1, lemma1_2]
addSEs [equals0D]) 1);
val lemma1 = result();
goal thy "!!A. 0 ~: A ==> 0 ~: {(nat -> Union(A)) * C. C:A}";
-by (fast_tac (ZF_cs addEs [RepFunE,
+by (fast_tac (!claset addEs [RepFunE,
Sigma_fun_space_not0 RS not_sym RS notE]) 1);
val lemma2 = result();
@@ -119,11 +117,11 @@
by (rtac allI 1);
by (rtac impI 1);
by (excluded_middle_tac "A=0" 1);
-by (fast_tac (ZF_cs addSIs [not_emptyI, empty_fun]) 2);
+by (fast_tac (!claset addSIs [not_emptyI, empty_fun]) 2);
by (rtac lemma1 1);
by (etac allE 1);
by (etac impE 1 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSEs [RepFunE]
+by (fast_tac (!claset addSEs [RepFunE]
addSIs [lemma2, all_eqpoll_imp_pair_eqpoll,
Sigma_fun_space_eqpoll]) 1);
qed "AC7_AC6";
@@ -141,14 +139,14 @@
by (dtac bspec 1 THEN (assume_tac 1));
by (REPEAT (eresolve_tac [exE,conjE] 1));
by (hyp_subst_tac 1);
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSEs [sym RS equals0D]) 1);
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSEs [sym RS equals0D]) 1);
val lemma1 = result();
goal thy "!!A. [| f: (PROD X:RepFun(A,p). X); D:A |] \
\ ==> (lam x:A. f`p(x))`D : p(D)";
by (resolve_tac [beta RS ssubst] 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [apply_type]) 1);
+by (fast_tac (!claset addSEs [apply_type]) 1);
val lemma2 = result();
goalw thy AC_defs "!!Z. AC1 ==> AC8";
@@ -157,7 +155,7 @@
by (rtac impI 1);
by (etac impE 1);
by (etac lemma1 1);
-by (fast_tac (AC_cs addSEs [lemma2]) 1);
+by (fast_tac (!claset addSEs [lemma2]) 1);
qed "AC1_AC8";
@@ -169,11 +167,11 @@
goal thy "!!A. ALL B1:A. ALL B2:A. B1 eqpoll B2 ==> \
\ ALL B:A*A. EX B1 B2. B=<B1,B2> & B1 eqpoll B2";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val lemma1 = result();
goal thy "!!f. f:bij(fst(<a,b>),snd(<a,b>)) ==> f:bij(a,b)";
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val lemma2 = result();
goalw thy AC_defs "!!Z. AC8 ==> AC9";
@@ -182,7 +180,7 @@
by (etac allE 1);
by (etac impE 1);
by (etac lemma1 1);
-by (fast_tac (AC_cs addSEs [lemma2]) 1);
+by (fast_tac (!claset addSEs [lemma2]) 1);
qed "AC8_AC9";
@@ -197,12 +195,7 @@
(* Rules nedded to prove lemma1 *)
val snd_lepoll_SigmaI = prod_lepoll_self RS
((prod_commute_eqpoll RS eqpoll_imp_lepoll) RSN (2,lepoll_trans));
-val lemma1_cs = FOL_cs addSEs [UnE, RepFunE]
- addSIs [all_eqpoll_imp_pair_eqpoll, ballI,
- nat_cons_eqpoll RS eqpoll_trans]
- addEs [Sigma_fun_space_not0 RS not_emptyE]
- addIs [snd_lepoll_SigmaI, eqpoll_refl RSN
- (2, prod_eqpoll_cong), Sigma_fun_space_eqpoll];
+
goal thy "!!A. [| 0~:A; A~=0 |] \
\ ==> ALL B1: ({((nat->Union(A))*B)*nat. B:A} \
@@ -210,7 +203,11 @@
\ ALL B2: ({((nat->Union(A))*B)*nat. B:A} \
\ Un {cons(0,((nat->Union(A))*B)*nat). B:A}). \
\ B1 eqpoll B2";
-by (fast_tac lemma1_cs 1);
+by (fast_tac (!claset addSIs [all_eqpoll_imp_pair_eqpoll, ballI,
+ nat_cons_eqpoll RS eqpoll_trans]
+ addEs [Sigma_fun_space_not0 RS not_emptyE]
+ addIs [snd_lepoll_SigmaI, eqpoll_refl RSN
+ (2, prod_eqpoll_cong), Sigma_fun_space_eqpoll]) 1);
val lemma1 = result();
goal thy "!!A. ALL B1:{(F*B)*N. B:A} Un {cons(0,(F*B)*N). B:A}. \
@@ -222,7 +219,7 @@
by (rtac snd_type 1);
by (rtac fst_type 1);
by (resolve_tac [consI1 RSN (2, apply_type)] 1);
-by (fast_tac (ZF_cs addSIs [fun_weaken_type, bij_is_fun]) 1);
+by (fast_tac (!claset addSIs [fun_weaken_type, bij_is_fun]) 1);
val lemma2 = result();
goalw thy AC_defs "!!Z. AC9 ==> AC1";
@@ -230,8 +227,8 @@
by (rtac impI 1);
by (etac allE 1);
by (excluded_middle_tac "A=0" 1);
-by (fast_tac (FOL_cs addSIs [empty_fun]) 2);
by (etac impE 1);
by (rtac lemma1 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac (AC_cs addSEs [lemma2]) 1);
+by (fast_tac (!claset addSEs [lemma2]) 1);
+by (fast_tac (!claset addSIs [empty_fun]) 1);
qed "AC9_AC1";
--- a/src/ZF/AC/AC_Equiv.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC_Equiv.ML Fri Jan 03 15:01:55 1997 +0100
@@ -16,9 +16,6 @@
val AC_aux_defs = [pairwise_disjoint_def, sets_of_size_between_def];
-val AC_cs = OrdQuant_cs;
-val AC_ss = OrdQuant_ss;
-
(* ******************************************** *)
(* Theorems analogous to ones presented in "ZF/Ordinal.ML" *)
@@ -27,12 +24,11 @@
val [prem] = goalw Cardinal.thy [lepoll_def]
"m:nat ==> ALL n: nat. m le n --> m lepoll n";
by (nat_ind_tac "m" [prem] 1);
-by (fast_tac (ZF_cs addSIs [le_imp_subset RS id_subset_inj]) 1);
+by (fast_tac (!claset addSIs [le_imp_subset RS id_subset_inj]) 1);
by (rtac ballI 1);
by (eres_inst_tac [("n","n")] natE 1);
-by (asm_simp_tac (ZF_ss addsimps [inj_def, succI1 RS Pi_empty2]) 1);
-by (fast_tac (ZF_cs addSDs [le0D]) 1);
-by (fast_tac (ZF_cs addSIs [le_imp_subset RS id_subset_inj]) 1);
+by (asm_simp_tac (!simpset addsimps [inj_def, succI1 RS Pi_empty2]) 1);
+by (fast_tac (!claset addSIs [le_imp_subset RS id_subset_inj]) 1);
val nat_le_imp_lepoll_lemma = result();
(* used in : AC10-AC15.ML WO1-WO6.ML WO6WO1.ML*)
@@ -43,14 +39,14 @@
(* ********************************************************************** *)
goal thy "!!X. (A->X)=0 ==> X=0";
-by (fast_tac (ZF_cs addSIs [equals0I] addEs [lam_type RSN (2, equals0D)]) 1);
+by (fast_tac (!claset addSIs [equals0I] addEs [lam_type RSN (2, equals0D)]) 1);
val fun_space_emptyD = result();
(* used only in WO1_DC.ML *)
(*Note simpler proof*)
-goal ZF.thy "!!A f g. [| ALL x:A. f`x=g`x; f:Df->Cf; g:Dg->Cg; \
+goal upair.thy "!!A f g. [| ALL x:A. f`x=g`x; f:Df->Cf; g:Dg->Cg; \
\ A<=Df; A<=Dg |] ==> f``A=g``A";
-by (asm_simp_tac (ZF_ss addsimps [image_fun]) 1);
+by (asm_simp_tac (!simpset addsimps [image_fun]) 1);
val images_eq = result();
(* used in : AC10-AC15.ML AC16WO4.ML WO6WO1.ML *)
@@ -63,7 +59,7 @@
by (res_inst_tac [("A2","A")]
(Diff_sing_lepoll RSN (2, subset_imp_lepoll RS lepoll_trans)) 1
THEN (REPEAT (assume_tac 2)));
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val Diff_lepoll = result();
(* ********************************************************************** *)
@@ -77,12 +73,12 @@
(* lemma for ordertype_Int *)
goalw Cardinal.thy [rvimage_def] "rvimage(A,id(A),r) = r Int A*A";
by (rtac equalityI 1);
-by (step_tac ZF_cs 1);
+by (Step_tac 1);
by (dres_inst_tac [("P","%a. <id(A)`xb,a>:r")] (id_conv RS subst) 1
THEN (assume_tac 1));
by (dres_inst_tac [("P","%a. <a,ya>:r")] (id_conv RS subst) 1
THEN (REPEAT (assume_tac 1)));
-by (fast_tac (ZF_cs addIs [id_conv RS ssubst]) 1);
+by (fast_tac (!claset addIs [id_conv RS ssubst]) 1);
val rvimage_id = result();
(* used only in Hartog.ML *)
@@ -96,42 +92,42 @@
(* used only in AC16_lemmas.ML *)
goalw CardinalArith.thy [InfCard_def]
"!!i. [| ~Finite(i); Card(i) |] ==> InfCard(i)";
-by (asm_simp_tac (ZF_ss addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1);
+by (asm_simp_tac (!simpset addsimps [Card_is_Ord RS nat_le_infinite_Ord]) 1);
val Inf_Card_is_InfCard = result();
goal thy "(THE z. {x}={z}) = x";
-by (fast_tac (AC_cs addSIs [the_equality]
+by (fast_tac (!claset addSIs [the_equality]
addSEs [singleton_eq_iff RS iffD1 RS sym]) 1);
val the_element = result();
goal thy "(lam x:A. {x}) : bij(A, {{x}. x:A})";
by (res_inst_tac [("d","%z. THE x. z={x}")] lam_bijective 1);
by (TRYALL (eresolve_tac [RepFunI, RepFunE]));
-by (REPEAT (asm_full_simp_tac (AC_ss addsimps [the_element]) 1));
+by (REPEAT (asm_full_simp_tac (!simpset addsimps [the_element]) 1));
val lam_sing_bij = result();
val [major,minor] = goal thy
"[| f : Pi(A,B); (!!x. x:A ==> B(x)<=C(x)) |] ==> f : Pi(A,C)";
-by (fast_tac (AC_cs addSIs [major RS Pi_type, minor RS subsetD,
+by (fast_tac (!claset addSIs [major RS Pi_type, minor RS subsetD,
major RS apply_type]) 1);
val Pi_weaken_type = result();
val [major, minor] = goalw thy [inj_def]
"[| f:inj(A, B); (!!a. a:A ==> f`a : C) |] ==> f:inj(A,C)";
-by (fast_tac (AC_cs addSEs [minor]
+by (fast_tac (!claset addSEs [minor]
addSIs [major RS CollectD1 RS Pi_type, major RS CollectD2]) 1);
val inj_strengthen_type = result();
goal thy "A*B=0 <-> A=0 | B=0";
-by (fast_tac (ZF_cs addSIs [equals0I] addEs [equals0D]) 1);
+by (fast_tac (!claset addSIs [equals0I] addEs [equals0D]) 1);
val Sigma_empty_iff = result();
goalw thy [Finite_def] "!!n. n:nat ==> Finite(n)";
-by (fast_tac (AC_cs addSIs [eqpoll_refl]) 1);
+by (fast_tac (!claset addSIs [eqpoll_refl]) 1);
val nat_into_Finite = result();
goalw thy [Finite_def] "~Finite(nat)";
-by (fast_tac (AC_cs addSDs [eqpoll_imp_lepoll]
+by (fast_tac (!claset addSDs [eqpoll_imp_lepoll]
addIs [Ord_nat RSN (2, ltI) RS lt_not_lepoll RS notE]) 1);
val nat_not_Finite = result();
@@ -144,8 +140,8 @@
goal thy "!!x. [| EX! x. P(x); P(x); P(y) |] ==> x=y";
by (etac ex1E 1);
by (res_inst_tac [("b","xa")] (sym RSN (2, trans)) 1);
-by (fast_tac AC_cs 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
+by (Fast_tac 1);
val ex1_two_eq = result();
(* ********************************************************************** *)
@@ -155,16 +151,16 @@
goalw thy [surj_def] "!!f. f : surj(A, B) ==> f``A = B";
by (etac CollectE 1);
by (resolve_tac [subset_refl RSN (2, image_fun) RS ssubst] 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [RepFunE, apply_type]
+by (fast_tac (!claset addSEs [RepFunE, apply_type]
addSIs [RepFunI] addIs [equalityI]) 1);
val surj_image_eq = result();
goal thy "!!y. succ(x) lepoll y ==> y ~= 0";
-by (fast_tac (ZF_cs addSDs [lepoll_0_is_0]) 1);
+by (fast_tac (!claset addSDs [lepoll_0_is_0]) 1);
val succ_lepoll_imp_not_empty = result();
goal thy "!!x. x eqpoll succ(n) ==> x ~= 0";
-by (fast_tac (AC_cs addSEs [eqpoll_sym RS eqpoll_0_is_0 RS succ_neq_0]) 1);
+by (fast_tac (!claset addSEs [eqpoll_sym RS eqpoll_0_is_0 RS succ_neq_0]) 1);
val eqpoll_succ_imp_not_empty = result();
--- a/src/ZF/AC/AC_Equiv.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/AC_Equiv.thy Fri Jan 03 15:01:55 1997 +0100
@@ -12,7 +12,7 @@
but slightly changed.
*)
-AC_Equiv = CardinalArith + Univ + OrdQuant +
+AC_Equiv = CardinalArith + Univ +
consts
--- a/src/ZF/AC/Cardinal_aux.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/Cardinal_aux.ML Fri Jan 03 15:01:55 1997 +0100
@@ -15,15 +15,15 @@
(* j=|A| *)
goal Cardinal.thy
"!!A. [| A lepoll i; Ord(i) |] ==> EX j. j le i & A eqpoll j";
-by (fast_tac (ZF_cs addIs [lepoll_cardinal_le, well_ord_Memrel,
- well_ord_cardinal_eqpoll RS eqpoll_sym]
- addDs [lepoll_well_ord]) 1);
+by (fast_tac (!claset addIs [lepoll_cardinal_le, well_ord_Memrel,
+ well_ord_cardinal_eqpoll RS eqpoll_sym]
+ addDs [lepoll_well_ord]) 1);
qed "lepoll_imp_ex_le_eqpoll";
(* j=|A| *)
goalw Cardinal.thy [lesspoll_def]
"!!A a. [| A lesspoll i; Ord(i) |] ==> EX j. j<i & A eqpoll j";
-by (fast_tac (ZF_cs addSDs [lepoll_imp_ex_le_eqpoll] addSEs [leE]) 1);
+by (fast_tac (!claset addSDs [lepoll_imp_ex_le_eqpoll] addSEs [leE]) 1);
qed "lesspoll_imp_ex_lt_eqpoll";
goalw thy [InfCard_def] "!!i. [| ~Finite(i); Ord(i) |] ==> InfCard(|i|)";
@@ -49,7 +49,7 @@
by (rtac eqpollI 1);
by (eresolve_tac [subset_imp_lepoll RSN (2, eqpoll_sym RS eqpoll_imp_lepoll
RS lepoll_trans)] 2);
-by (fast_tac AC_cs 2);
+by (Fast_tac 2);
by (resolve_tac [Un_lepoll_sum RS lepoll_trans] 1);
by (resolve_tac [lepoll_imp_sum_lepoll_prod RS lepoll_trans] 1);
by (eresolve_tac [eqpoll_sym RSN (2, eqpoll_trans) RS eqpoll_imp_lepoll] 1
@@ -65,33 +65,33 @@
THEN REPEAT (assume_tac 1));
val Un_eqpoll_Inf_Ord = result();
-val ss = ZF_ss addsimps [inj_is_fun RS apply_type, left_inverse]
+val ss = (!simpset) addsimps [inj_is_fun RS apply_type, left_inverse]
setloop (split_tac [expand_if] ORELSE' etac UnE);
-goal ZF.thy "{x, y} - {y} = {x} - {y}";
+goal upair.thy "{x, y} - {y} = {x} - {y}";
by (fast_tac eq_cs 1);
val double_Diff_sing = result();
-goal ZF.thy "if({y,z}-{z}=0, z, THE w. {y,z}-{z}={w}) = y";
+goal upair.thy "if({y,z}-{z}=0, z, THE w. {y,z}-{z}={w}) = y";
by (split_tac [expand_if] 1);
-by (asm_full_simp_tac (ZF_ss addsimps [double_Diff_sing, Diff_eq_0_iff]) 1);
-by (fast_tac (ZF_cs addSIs [the_equality, equalityI, equals0I]
+by (asm_full_simp_tac (!simpset addsimps [double_Diff_sing, Diff_eq_0_iff]) 1);
+by (fast_tac (!claset addSIs [the_equality, equalityI, equals0I]
addEs [equalityE]) 1);
val paired_bij_lemma = result();
goal thy "(lam y:{{y,z}. y:x}. if(y-{z}=0, z, THE w. y-{z}={w})) \
\ : bij({{y,z}. y:x}, x)";
by (res_inst_tac [("d","%a. {a,z}")] lam_bijective 1);
-by (TRYALL (fast_tac (AC_cs addSEs [RepFunE] addSIs [RepFunI]
- addss (AC_ss addsimps [paired_bij_lemma]))));
+by (TRYALL (fast_tac (!claset addSEs [RepFunE] addSIs [RepFunI]
+ addss (!simpset addsimps [paired_bij_lemma]))));
val paired_bij = result();
goalw thy [eqpoll_def] "{{y,z}. y:x} eqpoll x";
-by (fast_tac (AC_cs addSIs [paired_bij]) 1);
+by (fast_tac (!claset addSIs [paired_bij]) 1);
val paired_eqpoll = result();
goal thy "!!A. EX B. B eqpoll A & B Int C = 0";
-by (fast_tac (AC_cs addSIs [paired_eqpoll, equals0I] addEs [mem_asym]) 1);
+by (fast_tac (!claset addSIs [paired_eqpoll, equals0I] addEs [mem_asym]) 1);
val ex_eqpoll_disjoint = result();
goal thy "!!A. [| A lepoll i; B lepoll i; ~Finite(i); Ord(i) |] \
@@ -110,28 +110,28 @@
by (eresolve_tac [Least_le RS leE] 1);
by (etac Ord_in_Ord 1 THEN (assume_tac 1));
by (etac ltE 1);
-by (fast_tac (AC_cs addDs [OrdmemD]) 1);
+by (fast_tac (!claset addDs [OrdmemD]) 1);
by (etac subst_elem 1 THEN (assume_tac 1));
val Least_in_Ord = result();
goal thy "!!x. [| well_ord(x,r); y<=x; y lepoll succ(n); n:nat |] \
\ ==> y-{THE b. first(b,y,r)} lepoll n";
by (res_inst_tac [("Q","y=0")] (excluded_middle RS disjE) 1);
-by (fast_tac (AC_cs addSIs [Diff_sing_lepoll, the_first_in]) 1);
+by (fast_tac (!claset addSIs [Diff_sing_lepoll, the_first_in]) 1);
by (res_inst_tac [("b","y-{THE b. first(b, y, r)}")] subst 1);
by (rtac empty_lepollI 2);
-by (fast_tac (AC_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val Diff_first_lepoll = result();
goal thy "(UN x:X. P(x)) <= (UN x:X. P(x)-Q(x)) Un (UN x:X. Q(x))";
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val UN_subset_split = result();
goalw thy [lepoll_def] "!!a. Ord(a) ==> (UN x:a. {P(x)}) lepoll a";
by (res_inst_tac [("x","lam z:(UN x:a. {P(x)}). (LEAST i. P(i)=z)")] exI 1);
by (res_inst_tac [("d","%z. P(z)")] lam_injective 1);
-by (fast_tac (AC_cs addSIs [Least_in_Ord]) 1);
-by (fast_tac (AC_cs addIs [LeastI] addSEs [Ord_in_Ord]) 1);
+by (fast_tac (!claset addSIs [Least_in_Ord]) 1);
+by (fast_tac (!claset addIs [LeastI] addSEs [Ord_in_Ord]) 1);
val UN_sing_lepoll = result();
goal thy "!!a T. [| well_ord(T, R); ~Finite(a); Ord(a); n:nat |] ==> \
@@ -143,14 +143,14 @@
by (rtac empty_lepollI 2);
by (resolve_tac [equals0I RS sym] 1);
by (REPEAT (eresolve_tac [UN_E, allE] 1));
-by (fast_tac (AC_cs addDs [lepoll_0_is_0 RS subst]) 1);
+by (fast_tac (!claset addDs [lepoll_0_is_0 RS subst]) 1);
by (rtac allI 1);
by (rtac impI 1);
by (eres_inst_tac [("x","lam x:a. f`x - {THE b. first(b,f`x,R)}")] allE 1);
by (etac impE 1);
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSIs [Diff_first_lepoll]) 1);
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSIs [Diff_first_lepoll]) 1);
+by (Asm_full_simp_tac 1);
by (resolve_tac [UN_subset_split RS subset_imp_lepoll RS lepoll_trans] 1);
by (rtac Un_lepoll_Inf_Ord 1 THEN (REPEAT_FIRST assume_tac));
by (etac UN_sing_lepoll 1);
@@ -159,7 +159,7 @@
goal thy "!!a f. [| ALL b:a. f`b lepoll n & f`b <= T; well_ord(T, R); \
\ ~Finite(a); Ord(a); n:nat |] ==> (UN b:a. f`b) lepoll a";
by (eresolve_tac [UN_fun_lepoll_lemma RS allE] 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val UN_fun_lepoll = result();
goal thy "!!a f. [| ALL b:a. F(b) lepoll n & F(b) <= T; well_ord(T, R); \
@@ -167,13 +167,13 @@
by (rtac impE 1 THEN (assume_tac 3));
by (res_inst_tac [("f","lam b:a. F(b)")] (UN_fun_lepoll) 2
THEN (TRYALL assume_tac));
-by (simp_tac AC_ss 2);
-by (asm_full_simp_tac AC_ss 1);
+by (Simp_tac 2);
+by (Asm_full_simp_tac 1);
val UN_lepoll = result();
goal thy "!!a. Ord(a) ==> (UN b:a. F(b)) = (UN b:a. F(b) - (UN c:b. F(c)))";
by (rtac equalityI 1);
-by (fast_tac AC_cs 2);
+by (Fast_tac 2);
by (rtac subsetI 1);
by (etac UN_E 1);
by (rtac UN_I 1);
@@ -189,11 +189,11 @@
goalw thy [eqpoll_def] "!!A B. A Int B = 0 ==> A Un B eqpoll A + B";
by (res_inst_tac [("x","lam a:A Un B. if(a:A,Inl(a),Inr(a))")] exI 1);
by (res_inst_tac [("d","%z. case(%x.x, %x.x, z)")] lam_bijective 1);
-by (fast_tac (AC_cs addSIs [if_type, InlI, InrI]) 1);
+by (fast_tac (!claset addSIs [if_type, InlI, InrI]) 1);
by (TRYALL (etac sumE ));
by (TRYALL (split_tac [expand_if]));
-by (TRYALL (asm_simp_tac sum_ss));
-by (fast_tac (AC_cs addDs [equals0D]) 1);
+by (TRYALL Asm_simp_tac);
+by (fast_tac (!claset addDs [equals0D]) 1);
val disj_Un_eqpoll_sum = result();
goalw thy [lepoll_def, eqpoll_def]
@@ -201,7 +201,7 @@
by (etac exE 1);
by (forward_tac [subset_refl RSN (2, restrict_bij)] 1);
by (res_inst_tac [("x","f``a")] exI 1);
-by (fast_tac (AC_cs addSEs [inj_is_fun RS fun_is_rel RS image_subset]) 1);
+by (fast_tac (!claset addSEs [inj_is_fun RS fun_is_rel RS image_subset]) 1);
val lepoll_imp_eqpoll_subset = result();
(* ********************************************************************** *)
@@ -227,22 +227,22 @@
by (dresolve_tac [[lepoll_Finite, lepoll_Finite] MRS Finite_Un] 2
THEN (REPEAT (assume_tac 2)));
by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_Finite] 2);
-by (fast_tac (AC_cs
+by (fast_tac (!claset
addDs [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite]) 2);
by (dresolve_tac [ Un_lepoll_Inf_Ord] 1
THEN (REPEAT (assume_tac 1)));
-by (fast_tac (AC_cs addSEs [ltE, Ord_in_Ord]) 1);
+by (fast_tac (!claset addSEs [ltE, Ord_in_Ord]) 1);
by (dresolve_tac [subset_Un_Diff RS subset_imp_lepoll RS lepoll_trans RSN
(3, lt_Card_imp_lesspoll RS lepoll_lesspoll_lesspoll)] 1
THEN (TRYALL assume_tac));
-by (fast_tac (AC_cs addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1);
+by (fast_tac (!claset addSDs [lesspoll_def RS def_imp_iff RS iffD1]) 1);
val Diff_lesspoll_eqpoll_Card_lemma = result();
goal thy "!!A. [| A eqpoll a; ~Finite(a); Card(a); B lesspoll a |] \
\ ==> A - B eqpoll a";
-by (rtac swap 1 THEN (fast_tac AC_cs 1));
+by (rtac swap 1 THEN (Fast_tac 1));
by (rtac Diff_lesspoll_eqpoll_Card_lemma 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac (AC_cs addSIs [lesspoll_def RS def_imp_iff RS iffD2,
+by (fast_tac (!claset addSIs [lesspoll_def RS def_imp_iff RS iffD2,
subset_imp_lepoll RS (eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1);
val Diff_lesspoll_eqpoll_Card = result();
--- a/src/ZF/AC/Cardinal_aux.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/Cardinal_aux.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,3 @@
(*Dummy theory to document dependencies *)
-Cardinal_aux = AC_Equiv + first
+Cardinal_aux = AC_Equiv
--- a/src/ZF/AC/DC.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/DC.ML Fri Jan 03 15:01:55 1997 +0100
@@ -33,16 +33,16 @@
(* *)
(* ********************************************************************** *)
-goal thy "{z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)} <= XX*XX";
-by (fast_tac AC_cs 1);
+goal thy "{<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1} <= XX*XX";
+by (Fast_tac 1);
val lemma1_1 = result();
goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R) \
-\ ==> {z: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \
+\ ==> {<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \
\ (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}). \
-\ domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)} ~= 0";
+\ domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1} ~= 0";
by (etac ballE 1);
by (eresolve_tac [empty_subsetI RS PowI RSN (2, notE)] 2);
by (eresolve_tac [nat_0I RS n_lesspoll_nat RSN (2, impE)] 1);
@@ -50,23 +50,23 @@
by (res_inst_tac [("a","<0, {<0, x>}>")] not_emptyI 1);
by (rtac CollectI 1);
by (rtac SigmaI 1);
-by (fast_tac (AC_cs addSIs [nat_0I RS UN_I, empty_fun]) 1);
-by (fast_tac (AC_cs addSIs [nat_1I RS UN_I, singleton_fun RS Pi_type]
- addss (AC_ss addsimps [[lepoll_refl, succI1] MRS lepoll_1_is_sing,
- apply_singleton_eq, image_0])) 1);
-by (asm_full_simp_tac (AC_ss addsimps [domain_0, restrict_0, domain_cons,
- [lepoll_refl, succI1] MRS lepoll_1_is_sing]) 1);
+by (fast_tac (!claset addSIs [nat_0I RS UN_I, empty_fun]) 1);
+br (nat_1I RS UN_I) 1;
+by (fast_tac (!claset addSIs [singleton_fun RS Pi_type]
+ addss (!simpset addsimps [singleton_0 RS sym])) 1);
+by (asm_full_simp_tac (!simpset addsimps [domain_0, domain_cons,
+ singleton_0]) 1);
val lemma1_2 = result();
goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R) \
-\ ==> range({z: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \
+\ ==> range({<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \
\ (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}). \
-\ domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)}) \
-\ <= domain({z: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \
+\ domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1}) \
+\ <= domain({<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \
\ (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}). \
-\ domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)})";
+\ domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1})";
by (rtac range_subset_domain 1);
by (rtac subsetI 2);
by (etac CollectD1 2);
@@ -81,92 +81,82 @@
by (res_inst_tac [("x","cons(<n,x>, g)")] exI 1);
by (rtac CollectI 1);
by (rtac SigmaI 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (rtac UN_I 1);
by (etac nat_succI 1);
by (rtac CollectI 1);
by (etac cons_fun_type2 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [succE] addss (AC_ss
+by (fast_tac (!claset addSEs [succE] addss (!simpset
addsimps [cons_image_n, cons_val_n, cons_image_k, cons_val_k])) 1);
-by (asm_full_simp_tac (AC_ss
+by (asm_full_simp_tac (!simpset
addsimps [domain_cons, domain_of_fun, succ_def, restrict_cons_eq]) 1);
val lemma1_3 = result();
goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \
-\ RR = {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \
+\ RR = {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1}; \
\ ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R) \
\ |] ==> RR <= XX*XX & RR ~= 0 & range(RR) <= domain(RR)";
-by (fast_tac (AC_cs addSIs [lemma1_1] addSEs [lemma1_2, lemma1_3]) 1);
+by (fast_tac (!claset addSIs [lemma1_1] addSEs [lemma1_2, lemma1_3]) 1);
val lemma1 = result();
-goal thy "!!f. [| <f,g> : {z:XX*XX. \
-\ domain(snd(z))=succ(domain(fst(z))) & Q(z)}; \
-\ XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); f:k->X \
-\ |] ==> g:succ(k)->X";
-by (etac CollectE 1);
-by (dtac SigmaD2 1);
-by (hyp_subst_tac 1);
-by (etac UN_E 1);
-by (etac CollectE 1);
-by (asm_full_simp_tac AC_ss 1);
-by (dtac domain_of_fun 1);
-by (etac conjE 1);
-by (forward_tac [domain_of_fun RS sym RS trans] 1 THEN (assume_tac 1));
-by (fast_tac AC_cs 1);
-val lemma2_1 = result();
-
-goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \
+goal thy
+"!!X.[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \
\ ALL n:nat. <f`n, f`succ(n)> : \
-\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \
+\ {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1}; \
\ f: nat -> XX; n:nat \
-\ |] ==> EX k:nat. f`succ(n) : k -> X & n:k \
-\ & <f`succ(n)``n, f`succ(n)`n> : R";
+\ |] ==> EX k:nat. f`succ(n) : k -> X & n:k \
+\ & <f`succ(n)``n, f`succ(n)`n> : R";
by (etac nat_induct 1);
by (dresolve_tac [nat_1I RSN (2, apply_type)] 1);
by (dresolve_tac [nat_0I RSN (2, bspec)] 1);
-by (asm_full_simp_tac AC_ss 1);
-by (etac UN_E 1);
-by (etac CollectE 1);
+by (Asm_full_simp_tac 1);
+by (Step_tac 1);
by (rtac bexI 1 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSEs [nat_0_le RS leE, ltD, ltD RSN (2, bspec)]
- addEs [sym RS trans RS succ_neq_0, domain_of_fun]) 1);
-by (etac bexE 1);
+by (best_tac (!claset addIs [ltD]
+ addSEs [nat_0_le RS leE]
+ addEs [sym RS trans RS succ_neq_0, domain_of_fun]
+ addss (!simpset)) 1);
+(** LEVEL 7 **)
by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (etac conjE 1);
-by (dtac lemma2_1 1 THEN REPEAT (assume_tac 1));
-by (hyp_subst_tac 1);
+by (subgoal_tac "f ` succ(succ(x)) : succ(k)->X" 1);
by (dresolve_tac [nat_succI RS nat_succI RSN (2, apply_type)] 1
THEN (assume_tac 1));
-by (etac UN_E 1);
-by (etac CollectE 1);
-by (dresolve_tac [[domain_of_fun RS sym, domain_of_fun] MRS trans] 1
- THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [nat_succI, nat_into_Ord RS succ_in_succ]
+by (Full_simp_tac 1);
+by (Step_tac 1);
+by (forw_inst_tac [("a","succ(k)")] (domain_of_fun RS sym RS trans) 1 THEN
+ (assume_tac 1));
+by (forw_inst_tac [("a","xa")] (domain_of_fun RS sym RS trans) 1 THEN
+ (assume_tac 1));
+by (fast_tac (!claset addSEs [nat_succI, nat_into_Ord RS succ_in_succ]
addSDs [nat_into_Ord RS succ_in_succ RSN (2, bspec)]) 1);
+by (dtac domain_of_fun 1);
+by (Full_simp_tac 1);
+by (deepen_tac (!claset addDs [domain_of_fun RS sym RS trans]) 0 1);
val lemma2 = result();
-goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \
+goal thy
+"!!X.[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \
\ ALL n:nat. <f`n, f`succ(n)> : \
-\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \
+\ {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1}; \
\ f: nat -> XX; n:nat \
-\ |] ==> ALL x:n. f`succ(n)`x = f`succ(x)`x";
+\ |] ==> ALL x:n. f`succ(n)`x = f`succ(x)`x";
by (etac nat_induct 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (rtac ballI 1);
by (etac succE 1);
by (rtac restrict_eq_imp_val_eq 1);
by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
by (dtac lemma2 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSDs [domain_of_fun]) 1);
+by (fast_tac (!claset addSDs [domain_of_fun]) 1);
by (dres_inst_tac [("x","xa")] bspec 1 THEN (assume_tac 1));
by (eresolve_tac [sym RS trans RS sym] 1);
by (resolve_tac [restrict_eq_imp_val_eq RS sym] 1);
by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
by (dtac lemma2 1 THEN REPEAT (assume_tac 1));
by (fast_tac (FOL_cs addSDs [domain_of_fun]
addSEs [bexE, nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1);
@@ -174,31 +164,31 @@
goal thy "!!n. ALL x:n. f`succ(n)`x = f`succ(x)`x \
\ ==> {f`succ(x)`x. x:n} = {f`succ(n)`x. x:n}";
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val lemma3_2 = result();
goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \
\ ALL n:nat. <f`n, f`succ(n)> : \
-\ {z:XX*XX. domain(snd(z))=succ(domain(fst(z))) \
-\ & restrict(snd(z), domain(fst(z))) = fst(z)}; \
+\ {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \
+\ & restrict(z2, domain(z1)) = z1}; \
\ f: nat -> XX; n:nat \
\ |] ==> (lam x:nat. f`succ(x)`x) `` n = f`succ(n)``n";
by (etac natE 1);
-by (asm_full_simp_tac (AC_ss addsimps [image_0]) 1);
+by (asm_full_simp_tac (!simpset addsimps [image_0]) 1);
by (resolve_tac [image_lam RS ssubst] 1);
-by (fast_tac (AC_cs addSEs [[nat_succI, Ord_nat] MRS OrdmemD]) 1);
+by (fast_tac (!claset addSEs [[nat_succI, Ord_nat] MRS OrdmemD]) 1);
by (resolve_tac [lemma3_1 RS lemma3_2 RS ssubst] 1
THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSEs [nat_succI]) 1);
+by (fast_tac (!claset addSEs [nat_succI]) 1);
by (dresolve_tac [nat_succI RSN (4, lemma2)] 1
THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSEs [nat_into_Ord RSN (2, OrdmemD) RSN
- (2, image_fun RS sym)]) 1);
+by (fast_tac (!claset addSEs [nat_into_Ord RSN (2, OrdmemD) RSN
+ (2, image_fun RS sym)]) 1);
val lemma3 = result();
goal thy "!!f. [| f:A->B; B<=C |] ==> f:A->C";
by (rtac Pi_type 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [apply_type]) 1);
+by (fast_tac (!claset addSEs [apply_type]) 1);
val fun_type_gen = result();
goalw thy [DC_def, DC0_def] "!!Z. DC0 ==> DC(nat)";
@@ -208,16 +198,16 @@
THEN (assume_tac 1));
by (etac bexE 1);
by (res_inst_tac [("x","lam n:nat. f`succ(n)`n")] bexI 1);
-by (fast_tac (AC_cs addSIs [lam_type] addSDs [refl RS lemma2]
+by (fast_tac (!claset addSIs [lam_type] addSDs [refl RS lemma2]
addSEs [fun_type_gen, apply_type]) 2);
by (rtac oallI 1);
by (forward_tac [ltD RSN (3, refl RS lemma2)] 1
THEN assume_tac 2);
-by (fast_tac (AC_cs addSEs [fun_type_gen]) 1);
+by (fast_tac (!claset addSEs [fun_type_gen]) 1);
by (eresolve_tac [ltD RSN (3, refl RS lemma3) RS ssubst] 1
THEN assume_tac 2);
-by (fast_tac (AC_cs addSEs [fun_type_gen]) 1);
-by (fast_tac (AC_cs addss AC_ss) 1);
+by (fast_tac (!claset addSEs [fun_type_gen]) 1);
+by (fast_tac (!claset addss (!simpset)) 1);
qed "DC0_DC_nat";
(* ********************************************************************** *)
@@ -231,11 +221,11 @@
(* *)
(* XX = (UN n:nat. *)
(* {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}) *)
-(* RR = {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) *)
-(* & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | *)
-(* (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) *)
-(* & (ALL f:fst(z). restrict(g, domain(f)) = f)) & *)
-(* snd(z)={<0,x>})} *)
+(* RR = {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) *)
+(* & (ALL f:z1. restrict(z2, domain(f)) = f)) | *)
+(* (~ (EX g:XX. domain(g)=succ(UN f:z1. domain(f)) *)
+(* & (ALL f:z1. restrict(g, domain(f)) = f)) & *)
+(* z2={<0,x>})} *)
(* *)
(* Then XX and RR satisfy the hypotheses of DC(omega). *)
(* So applying DC: *)
@@ -252,14 +242,14 @@
goalw thy [lesspoll_def, Finite_def]
"!!A. A lesspoll nat ==> Finite(A)";
-by (fast_tac (AC_cs addSDs [ltD, lepoll_imp_ex_le_eqpoll]
+by (fast_tac (!claset addSDs [ltD, lepoll_imp_ex_le_eqpoll]
addSIs [Ord_nat]) 1);
val lesspoll_nat_is_Finite = result();
goal thy "!!n. n:nat ==> ALL A. (A eqpoll n & A <= X) --> A : Fin(X)";
by (etac nat_induct 1);
by (rtac allI 1);
-by (fast_tac (AC_cs addSIs [Fin.emptyI]
+by (fast_tac (!claset addSIs [Fin.emptyI]
addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1);
by (rtac allI 1);
by (rtac impI 1);
@@ -269,10 +259,10 @@
by (forward_tac [Diff_sing_eqpoll] 1 THEN (assume_tac 1));
by (etac allE 1);
by (etac impE 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (dtac subsetD 1 THEN (assume_tac 1));
by (dresolve_tac [Fin.consI] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (AC_ss addsimps [cons_Diff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [cons_Diff]) 1);
val Finite_Fin_lemma = result();
goalw thy [Finite_def] "!!A. [| Finite(A); A <= X |] ==> A : Fin(X)";
@@ -281,60 +271,62 @@
by (etac allE 1);
by (etac impE 1);
by (assume_tac 2);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val Finite_Fin = result();
-goal thy "!!x. x: X ==> {<0,x>}: (UN n:nat. \
-\ {f:succ(n)->X. ALL k:n. <f`k, f`succ(k)> : R})";
-by (fast_tac (AC_cs addSIs [nat_0I RS UN_I, singleton_fun RS Pi_type]
- addss (AC_ss addsimps [[lepoll_refl, succI1] MRS lepoll_1_is_sing,
- apply_singleton_eq])) 1);
+goal thy "!!x. x: X \
+\ ==> {<0,x>}: (UN n:nat. {f:succ(n)->X. ALL k:n. <f`k, f`succ(k)> : R})";
+br (nat_0I RS UN_I) 1;
+by (fast_tac (!claset addSIs [singleton_fun RS Pi_type]
+ addss (!simpset addsimps [singleton_0 RS sym])) 1);
val singleton_in_funs = result();
goal thy
"!!X. [| XX = (UN n:nat. \
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \
-\ RR = {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \
-\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | \
-\ (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) \
-\ & (ALL f:fst(z). restrict(g, domain(f)) = f)) & snd(z)={<0,x>})}; \
+\ RR = {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \
+\ & (ALL f:z1. restrict(z2, domain(f)) = f)) | \
+\ (~ (EX g:XX. domain(g)=succ(UN f:z1. domain(f)) \
+\ & (ALL f:z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}; \
\ range(R) <= domain(R); x:domain(R) \
\ |] ==> RR <= Pow(XX)*XX & \
\ (ALL Y:Pow(XX). Y lesspoll nat --> (EX x:XX. <Y,x>:RR))";
by (rtac conjI 1);
-by (deepen_tac (ZF_cs addSEs [FinD RS PowI]) 0 1);
+by (deepen_tac (!claset addSEs [FinD RS PowI]) 0 1);
by (rtac ballI 1);
by (rtac impI 1);
by (dresolve_tac [[lesspoll_nat_is_Finite, PowD] MRS Finite_Fin] 1
THEN (assume_tac 1));
by (excluded_middle_tac "EX g:XX. domain(g)=succ(UN f:Y. domain(f)) \
\ & (ALL f:Y. restrict(g, domain(f)) = f)" 1);
-by (fast_tac (AC_cs addss AC_ss) 2);
-by (fast_tac (FOL_cs addSEs [CollectE, singleton_in_funs]
- addSIs [SigmaI] addIs [bexI] addss AC_ss) 1);
+by (fast_tac (!claset addss (!simpset)) 2);
+by (step_tac (!claset delrules [domainE]) 1);
+by(swap_res_tac [bexI] 1 THEN etac singleton_in_funs 2);
+by (asm_simp_tac (!simpset addsimps [nat_0I RSN (2, bexI),
+ cons_fun_type2, empty_fun]) 1);
val lemma1 = result();
goal thy "!!f. [| f:nat->X; n:nat |] ==> \
\ (UN x:f``succ(n). P(x)) = P(f`n) Un (UN x:f``n. P(x))";
-by (asm_full_simp_tac (AC_ss
+by (asm_full_simp_tac (!simpset
addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun),
[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1);
val UN_image_succ_eq = result();
goal thy "!!f. [| (UN x:f``n. P(x)) = y; P(f`n) = succ(y); \
\ f:nat -> X; n:nat |] ==> (UN x:f``succ(n). P(x)) = succ(y)";
-by (asm_full_simp_tac (AC_ss addsimps [UN_image_succ_eq]) 1);
-by (fast_tac (AC_cs addSIs [equalityI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [UN_image_succ_eq]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val UN_image_succ_eq_succ = result();
-goal thy "!!f. [| f: (UN n:nat. {g:succ(n) -> D. P(g, n)}); \
+goal thy "!!f. [| f:succ(n) -> D; n:nat; \
\ domain(f)=succ(x); x=y |] ==> f`y : D";
-by (fast_tac (AC_cs addEs [apply_type]
+by (fast_tac (!claset addEs [apply_type]
addSDs [[sym, domain_of_fun] MRS trans]) 1);
-val apply_UN_type = result();
+val apply_domain_type = result();
goal thy "!!f. [| f : nat -> X; n:nat |] ==> f``succ(n) = cons(f`n, f``n)";
-by (asm_full_simp_tac (AC_ss
+by (asm_full_simp_tac (!simpset
addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1);
val image_fun_succ = result();
@@ -343,7 +335,7 @@
\ u=k; n:nat \
\ |] ==> f`n : succ(k) -> domain(R)";
by (dtac apply_type 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addEs [UN_E, domain_eq_imp_fun_type]) 1);
+by (fast_tac (!claset addEs [UN_E, domain_eq_imp_fun_type]) 1);
val f_n_type = result();
goal thy "!!f. [| f : nat -> (UN n:nat. \
@@ -354,8 +346,7 @@
by (etac UN_E 1);
by (etac CollectE 1);
by (dresolve_tac [domain_of_fun RS sym RS trans] 1 THEN (assume_tac 1));
-by (dtac succ_eqD 1);
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val f_n_pairs_in_R = result();
goalw thy [restrict_def]
@@ -363,9 +354,9 @@
\ |] ==> restrict(cons(<n, y>, f), domain(x)) = x";
by (eresolve_tac [sym RS trans RS sym] 1);
by (rtac fun_extension 1);
-by (fast_tac (AC_cs addSIs [lam_type]) 1);
-by (fast_tac (AC_cs addSIs [lam_type]) 1);
-by (asm_full_simp_tac (AC_ss addsimps [subsetD RS cons_val_k]) 1);
+by (fast_tac (!claset addSIs [lam_type]) 1);
+by (fast_tac (!claset addSIs [lam_type]) 1);
+by (asm_full_simp_tac (!simpset addsimps [subsetD RS cons_val_k]) 1);
val restrict_cons_eq_restrict = result();
goal thy "!!f. [| ALL x:f``n. restrict(f`n, domain(x))=x; \
@@ -375,45 +366,50 @@
\ (UN x:f``n. domain(x)) <= n |] \
\ ==> ALL x:f``succ(n). restrict(cons(<succ(n),y>, f`n), domain(x))=x";
by (rtac ballI 1);
-by (asm_full_simp_tac (AC_ss addsimps [image_fun_succ]) 1);
+by (asm_full_simp_tac (!simpset addsimps [image_fun_succ]) 1);
by (dtac f_n_type 1 THEN REPEAT (ares_tac [refl] 1));
-by (etac consE 1);
-by (asm_full_simp_tac (AC_ss addsimps [domain_of_fun, restrict_cons_eq]) 1);
+by (etac disjE 1);
+by (asm_full_simp_tac (!simpset addsimps [domain_of_fun, restrict_cons_eq]) 1);
by (dtac bspec 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [restrict_cons_eq_restrict]) 1);
+by (fast_tac (!claset addSEs [restrict_cons_eq_restrict]) 1);
val all_in_image_restrict_eq = result();
-goal thy "!!X. [| ALL b<nat. <f``b, f`b> : \
-\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \
-\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f)) | \
-\ (~ (EX g:XX. domain(g)=succ(UN f:fst(z). domain(f)) \
-\ & (ALL f:fst(z). restrict(g, domain(f)) = f)) & snd(z)={<0,x>})}; \
+goal thy
+"!!X.[| ALL b<nat. <f``b, f`b> : \
+\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \
+\ & (ALL f:z1. restrict(z2, domain(f)) = f)) | \
+\ (~ (EX g:XX. domain(g)=succ(UN f:z1. domain(f)) \
+\ & (ALL f:z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}; \
\ XX = (UN n:nat. \
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \
\ f: nat -> XX; range(R) <= domain(R); x:domain(R) \
-\ |] ==> ALL b<nat. <f``b, f`b> : \
-\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \
-\ & (UN f:fst(z). domain(f)) = b \
-\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}";
+\ |] ==> ALL b<nat. <f``b, f`b> : \
+\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \
+\ & (UN f:z1. domain(f)) = b \
+\ & (ALL f:z1. restrict(z2, domain(f)) = f))}";
by (rtac oallI 1);
by (dtac ltD 1);
by (etac nat_induct 1);
by (dresolve_tac [[nat_0I, Ord_nat] MRS ltI RSN (2, ospec)] 1);
by (fast_tac (FOL_cs addss
- (ZF_ss addsimps [image_0, singleton_fun RS domain_of_fun,
- [lepoll_refl, succI1] MRS lepoll_1_is_sing, singleton_in_funs])) 1);
+ (!simpset addsimps [image_0, singleton_fun RS domain_of_fun,
+ singleton_0, singleton_in_funs])) 1);
+(*induction step*) (** LEVEL 5 **)
+by (Full_simp_tac 1);
by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1
THEN (assume_tac 1));
-by (REPEAT (eresolve_tac [conjE, CollectE, disjE] 1));
+by (REPEAT (eresolve_tac [conjE, disjE] 1));
by (fast_tac (FOL_cs addSEs [trans, subst_context]
- addSIs [UN_image_succ_eq_succ] addss AC_ss) 1);
+ addSIs [UN_image_succ_eq_succ] addss (!simpset)) 1);
by (etac conjE 1);
by (etac notE 1);
-by (asm_full_simp_tac (AC_ss addsimps [UN_image_succ_eq_succ]) 1);
-by (etac conjE 1);
-by (dtac apply_UN_type 1 THEN REPEAT (assume_tac 1));
+by (asm_full_simp_tac (!simpset addsimps [UN_image_succ_eq_succ]) 1);
+(** LEVEL 12 **)
+by (REPEAT (eresolve_tac [conjE, bexE] 1));
+by (dtac apply_domain_type 1 THEN REPEAT (assume_tac 1));
by (etac domainE 1);
by (etac domainE 1);
+
by (forward_tac [f_n_type] 1 THEN REPEAT (assume_tac 1));
by (res_inst_tac [("x","cons(<succ(xa), ya>, f`xa)")] bexI 1);
by (fast_tac (FOL_cs
@@ -426,52 +422,53 @@
THEN REPEAT (assume_tac 1));
by (rtac ballI 1);
by (etac succE 1);
-by (asm_full_simp_tac (AC_ss addsimps [cons_val_n, cons_val_k]) 1);
-by (dresolve_tac [domain_of_fun RSN (2, f_n_pairs_in_R)] 1
- THEN REPEAT (assume_tac 1));
-by (dtac bspec 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (AC_ss
- addsimps [nat_into_Ord RS succ_in_succ, succI2, cons_val_k]) 1);
+(** LEVEL 25 **)
+by (dresolve_tac [domain_of_fun RSN (2, f_n_pairs_in_R)] 2
+ THEN REPEAT (assume_tac 2));
+by (dtac bspec 2 THEN (assume_tac 2));
+by (asm_full_simp_tac (!simpset
+ addsimps [nat_into_Ord RS succ_in_succ, succI2, cons_val_k]) 2);
+by (asm_full_simp_tac (!simpset addsimps [cons_val_n, cons_val_k]) 1);
val simplify_recursion = result();
+
goal thy "!!X. [| XX = (UN n:nat. \
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \
\ ALL b<nat. <f``b, f`b> : \
-\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \
-\ & (UN f:fst(z). domain(f)) = b \
-\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}; \
+\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \
+\ & (UN f:z1. domain(f)) = b \
+\ & (ALL f:z1. restrict(z2, domain(f)) = f))}; \
\ f: nat -> XX; range(R) <= domain(R); x:domain(R); n:nat \
\ |] ==> f`n : succ(n) -> domain(R) \
\ & (ALL i:n. <f`n`i, f`n`succ(i)>:R)";
by (dtac ospec 1);
by (eresolve_tac [Ord_nat RSN (2, ltI)] 1);
by (etac CollectE 1);
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
by (rtac conjI 1);
-by (fast_tac (AC_cs
+by (fast_tac (!claset
addSEs [trans RS domain_eq_imp_fun_type, subst_context]) 1);
-by (fast_tac (FOL_cs
- addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 1);
+by (fast_tac (FOL_cs addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 1);
val lemma2 = result();
goal thy "!!n. [| XX = (UN n:nat. \
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \
\ ALL b<nat. <f``b, f`b> : \
-\ {z:Fin(XX)*XX. (domain(snd(z))=succ(UN f:fst(z). domain(f)) \
-\ & (UN f:fst(z). domain(f)) = b \
-\ & (ALL f:fst(z). restrict(snd(z), domain(f)) = f))}; \
+\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \
+\ & (UN f:z1. domain(f)) = b \
+\ & (ALL f:z1. restrict(z2, domain(f)) = f))}; \
\ f : nat -> XX; n:nat; range(R) <= domain(R); x:domain(R) \
\ |] ==> f`n`n = f`succ(n)`n";
by (forward_tac [lemma2 RS conjunct1 RS domain_of_fun] 1
THEN REPEAT (assume_tac 1));
by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1
THEN (assume_tac 1));
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
by (REPEAT (etac conjE 1));
by (etac ballE 1);
by (eresolve_tac [restrict_eq_imp_val_eq RS sym] 1);
-by (fast_tac (AC_cs addSEs [ssubst]) 1);
-by (asm_full_simp_tac (AC_ss
+by (fast_tac (!claset addSEs [ssubst]) 1);
+by (asm_full_simp_tac (!simpset
addsimps [[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1);
val lemma3 = result();
@@ -491,7 +488,7 @@
by (forward_tac [refl RS (nat_succI RSN (6, lemma2)) RS conjunct2] 1
THEN REPEAT (assume_tac 1));
by (dresolve_tac [refl RS lemma3] 1 THEN REPEAT (assume_tac 1));
-by (asm_full_simp_tac (AC_ss addsimps [nat_succI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [nat_succI]) 1);
qed "DC_nat_DC0";
(* ********************************************************************** *)
@@ -500,7 +497,7 @@
goalw thy [lesspoll_def]
"!!A. [| ~ A lesspoll B; C lesspoll B |] ==> A - C ~= 0";
-by (fast_tac (AC_cs addSDs [Diff_eq_0_iff RS iffD1 RS subset_imp_lepoll]
+by (fast_tac (!claset addSDs [Diff_eq_0_iff RS iffD1 RS subset_imp_lepoll]
addSIs [eqpollI] addEs [notE] addSEs [eqpollE, lepoll_trans]) 1);
val lemma1 = result();
@@ -512,27 +509,27 @@
by (resolve_tac [Ord_a RS Ord_in_Ord RS Ord_linear_lt] 1
THEN (assume_tac 1));
by (eres_inst_tac [("j","x")] (Ord_a RS Ord_in_Ord) 1);
-by (REPEAT (fast_tac (AC_cs addDs [not_eq, not_eq RS not_sym]) 1));
+by (REPEAT (fast_tac (!claset addDs [not_eq, not_eq RS not_sym]) 1));
val fun_Ord_inj = result();
goal thy "!!a. [| f:X->Y; A<=X; a:A |] ==> f`a : f``A";
-by (fast_tac (AC_cs addSEs [image_fun RS ssubst]) 1);
+by (fast_tac (!claset addSEs [image_fun RS ssubst]) 1);
val value_in_image = result();
goalw thy [DC_def, WO3_def]
"!!Z. ALL K. Card(K) --> DC(K) ==> WO3";
by (rtac allI 1);
by (excluded_middle_tac "A lesspoll Hartog(A)" 1);
-by (fast_tac (AC_cs addSDs [lesspoll_imp_ex_lt_eqpoll]
+by (fast_tac (!claset addSDs [lesspoll_imp_ex_lt_eqpoll]
addSIs [Ord_Hartog, leI RS le_imp_subset]) 2);
by (REPEAT (eresolve_tac [allE, impE] 1));
by (rtac Card_Hartog 1);
by (eres_inst_tac [("x","A")] allE 1);
-by (eres_inst_tac [("x","{z:Pow(A)*A . fst(z) \
-\ lesspoll Hartog(A) & snd(z) ~: fst(z)}")] allE 1);
-by (asm_full_simp_tac AC_ss 1);
+by (eres_inst_tac [("x","{<z1,z2>:Pow(A)*A . z1 \
+\ lesspoll Hartog(A) & z2 ~: z1}")] allE 1);
+by (Asm_full_simp_tac 1);
by (etac impE 1);
-by (fast_tac (AC_cs addEs [lemma1 RS not_emptyE]) 1);
+by (fast_tac (!claset addEs [lemma1 RS not_emptyE]) 1);
by (etac bexE 1);
by (resolve_tac [exI RS (lepoll_def RS (def_imp_iff RS iffD2))
RS (HartogI RS notE)] 1);
@@ -540,7 +537,7 @@
by (dresolve_tac [Ord_Hartog RSN (2, OrdmemD) RSN (2,
ltD RSN (3, value_in_image))] 1
THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSDs [Ord_Hartog RSN (2, ltI) RSN (2, ospec)]
+by (fast_tac (!claset addSDs [Ord_Hartog RSN (2, ltI) RSN (2, ospec)]
addEs [subst]) 1);
qed "DC_WO3";
@@ -551,21 +548,21 @@
goal thy
"!!a. [| Ord(a); b:a |] ==> (lam x:a. P(x))``b = (lam x:b. P(x))``b";
by (rtac images_eq 1);
-by (REPEAT (fast_tac (AC_cs addSEs [RepFunI, OrdmemD]
+by (REPEAT (fast_tac (!claset addSEs [RepFunI, OrdmemD]
addSIs [lam_type]) 2));
by (rtac ballI 1);
by (dresolve_tac [OrdmemD RS subsetD] 1
THEN REPEAT (assume_tac 1));
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val lam_images_eq = result();
goalw thy [lesspoll_def] "!!K. [| Card(K); b:K |] ==> b lesspoll K";
-by (asm_full_simp_tac (AC_ss addsimps [Card_iff_initial]) 1);
-by (fast_tac (AC_cs addSIs [le_imp_lepoll, ltI, leI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [Card_iff_initial]) 1);
+by (fast_tac (!claset addSIs [le_imp_lepoll, ltI, leI]) 1);
val in_Card_imp_lesspoll = result();
goal thy "(lam b:a. P(b)) : a -> {P(b). b:a}";
-by (fast_tac (AC_cs addSIs [lam_type, RepFunI]) 1);
+by (fast_tac (!claset addSIs [lam_type, RepFunI]) 1);
val lam_type_RepFun = result();
goal thy "!!Z. [| ALL Y:Pow(X). Y lesspoll a --> (EX x:X. <Y, x> : R); \
@@ -584,13 +581,13 @@
by (rtac impI 1);
by (resolve_tac [ff_def RS def_transrec RS ssubst] 1);
by (etac the_first_in 1);
-by (fast_tac AC_cs 1);
-by (asm_full_simp_tac (AC_ss
+by (Fast_tac 1);
+by (asm_full_simp_tac (!simpset
addsimps [[lam_type_RepFun, subset_refl] MRS image_fun]) 1);
by (etac lemma_ 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [RepFunE, impE, notE]
+by (fast_tac (!claset addSEs [RepFunE, impE, notE]
addEs [Card_is_Ord RSN (2, OrdmemD) RS subsetD]) 1);
-by (fast_tac (AC_cs addSEs [[in_Card_imp_lesspoll, RepFun_lepoll]
+by (fast_tac (!claset addSEs [[in_Card_imp_lesspoll, RepFun_lepoll]
MRS lepoll_lesspoll_lesspoll]) 1);
val lemma = result();
@@ -601,8 +598,8 @@
by (res_inst_tac [("x","lam b:K. ff(b, X, Ra, R)")] bexI 1);
by (rtac lam_type 2);
by (resolve_tac [lemma RS CollectD1] 2 THEN REPEAT (assume_tac 2));
-by (asm_full_simp_tac (AC_ss
+by (asm_full_simp_tac (!simpset
addsimps [[Card_is_Ord, ltD] MRS lam_images_eq]) 1);
-by (fast_tac (AC_cs addSEs [ltE, lemma RS CollectD2]) 1);
+by (fast_tac (!claset addSEs [ltE, lemma RS CollectD2]) 1);
qed" WO1_DC_Card";
--- a/src/ZF/AC/DC.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/DC.thy Fri Jan 03 15:01:55 1997 +0100
@@ -5,7 +5,7 @@
Theory file for the proofs concernind the Axiom of Dependent Choice
*)
-DC = AC_Equiv + Hartog + first + Cardinal_aux + "DC_lemmas" +
+DC = AC_Equiv + Hartog + Cardinal_aux + "DC_lemmas" +
consts
--- a/src/ZF/AC/DC_lemmas.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/DC_lemmas.ML Fri Jan 03 15:01:55 1997 +0100
@@ -10,7 +10,7 @@
"Ord(a) ==> {P(b). b:a} lepoll a";
by (res_inst_tac [("x","lam z:RepFun(a,P). LEAST i. z=P(i)")] exI 1);
by (res_inst_tac [("d","%z. P(z)")] (sym RSN (2, lam_injective)) 1);
-by (fast_tac (AC_cs addSEs [RepFunE] addSIs [Least_in_Ord, prem]) 1);
+by (fast_tac (!claset addSEs [RepFunE] addSIs [Least_in_Ord, prem]) 1);
by (REPEAT (eresolve_tac [RepFunE, LeastI, prem RS Ord_in_Ord] 1));
val RepFun_lepoll = result();
@@ -29,73 +29,49 @@
"!!f. [| f:X->Y; Ord(X) |] ==> f``X lepoll X";
by (res_inst_tac [("x","lam x:f``X. LEAST y. f`y = x")] exI 1);
by (res_inst_tac [("d","%z. f`z")] lam_injective 1);
-by (fast_tac (AC_cs addSIs [Least_in_Ord, apply_equality]) 1);
-by (fast_tac (AC_cs addSEs [Ord_in_Ord] addSIs [LeastI, apply_equality]) 1);
+by (fast_tac (!claset addSIs [Least_in_Ord, apply_equality]) 1);
+by (fast_tac (!claset addSEs [Ord_in_Ord] addSIs [LeastI, apply_equality]) 1);
val image_Ord_lepoll = result();
-val apply_singleton_eq = [singletonI, singleton_fun] MRS apply_equality;
-
-goal thy "restrict(f, 0) = 0";
-by (resolve_tac [singleton_iff RS iffD1] 1);
-by (resolve_tac [Pi_empty1 RS subst] 1);
-by (fast_tac (AC_cs addSIs [restrict_type]) 1);
-val restrict_0 = result();
-
val [major, minor] = goal thy
"[| (!!g. g:X ==> EX u. <g,u>:R); R<=X*X \
\ |] ==> range(R) <= domain(R)";
by (rtac subsetI 1);
by (etac rangeE 1);
by (dresolve_tac [minor RS subsetD RS SigmaD2 RS major] 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val range_subset_domain = result();
val prems = goal thy "!!k. k:n ==> k~=n";
-by (fast_tac (AC_cs addSEs [mem_irrefl]) 1);
+by (fast_tac (!claset addSEs [mem_irrefl]) 1);
val mem_not_eq = result();
-goal thy "!!a. [| a:B; A<=B |] ==> cons(a, A) <= B";
-by (fast_tac AC_cs 1);
-val cons_subset = result();
-
-goal thy "!!g. g:n->X ==> cons(<n,x>, g) : succ(n) -> cons(x, X)";
-by (resolve_tac [Pi_iff_old RS iffD2] 1);
-by (rtac conjI 1);
-by (fast_tac (AC_cs addSIs [cons_subset, succI1]
- addEs [fun_is_rel RS subset_trans RS subsetD, succI2]) 1);
-by (rtac ballI 1);
-by (etac succE 1);
-by (fast_tac (AC_cs addDs [domain_of_fun, domainI] addSEs [mem_irrefl]) 1);
-by (fast_tac (AC_cs addDs [fun_unique_Pair] addSEs [mem_irrefl]) 1);
+goalw thy [succ_def] "!!g. g:n->X ==> cons(<n,x>, g) : succ(n) -> cons(x, X)";
+by (fast_tac (!claset addSIs [fun_extend] addSEs [mem_irrefl]) 1);
val cons_fun_type = result();
-goal thy "!!x. x:X ==> cons(x, X) = X";
-by (fast_tac (AC_cs addSIs [equalityI]) 1);
-val cons_eq_self = result();
-
-val [g_type, x_in] = goal thy "[| g:n->X; x:X |] ==> cons(<n,x>, g) : succ(n) -> X";
-by (resolve_tac [x_in RS cons_eq_self RS subst] 1);
-by (resolve_tac [g_type RS cons_fun_type] 1);
+goal thy "!!g. [| g:n->X; x:X |] ==> cons(<n,x>, g) : succ(n) -> X";
+by (etac (cons_absorb RS subst) 1 THEN etac cons_fun_type 1);
val cons_fun_type2 = result();
goal thy "!!n. n: nat ==> cons(<n,x>, g)``n = g``n";
-by (fast_tac (AC_cs addSIs [equalityI] addSEs [mem_irrefl]) 1);
+by (fast_tac (!claset addSIs [equalityI] addSEs [mem_irrefl]) 1);
val cons_image_n = result();
goal thy "!!n. g:n->X ==> cons(<n,x>, g)`n = x";
-by (fast_tac (AC_cs addSIs [apply_equality] addSEs [cons_fun_type]) 1);
+by (fast_tac (!claset addSIs [apply_equality] addSEs [cons_fun_type]) 1);
val cons_val_n = result();
goal thy "!!k. k : n ==> cons(<n,x>, g)``k = g``k";
-by (fast_tac (AC_cs addSIs [equalityI] addEs [mem_asym]) 1);
+by (fast_tac (!claset addSIs [equalityI] addEs [mem_asym]) 1);
val cons_image_k = result();
goal thy "!!k. [| k:n; g:n->X |] ==> cons(<n,x>, g)`k = g`k";
-by (fast_tac (AC_cs addSIs [apply_equality, consI2] addSEs [cons_fun_type, apply_Pair]) 1);
+by (fast_tac (!claset addSIs [apply_equality, consI2] addSEs [cons_fun_type, apply_Pair]) 1);
val cons_val_k = result();
goal thy "!!f. domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)";
-by (asm_full_simp_tac (AC_ss addsimps [domain_cons, succ_def]) 1);
+by (asm_full_simp_tac (!simpset addsimps [domain_cons, succ_def]) 1);
val domain_cons_eq_succ = result();
goalw thy [restrict_def] "!!g. g:n->X ==> restrict(cons(<n,x>, g), n)=g";
@@ -104,12 +80,12 @@
by (eresolve_tac [cons_fun_type RS apply_type] 1);
by (etac succI2 1);
by (assume_tac 1);
-by (asm_full_simp_tac (AC_ss addsimps [cons_val_k]) 1);
+by (asm_full_simp_tac (!simpset addsimps [cons_val_k]) 1);
val restrict_cons_eq = result();
goal thy "!!k. [| Ord(k); i:k |] ==> succ(i) : succ(k)";
by (resolve_tac [Ord_linear RS disjE] 1 THEN (assume_tac 3));
-by (REPEAT (fast_tac (AC_cs addSIs [Ord_succ]
+by (REPEAT (fast_tac (!claset addSIs [Ord_succ]
addEs [Ord_in_Ord, mem_irrefl, mem_asym]
addSDs [succ_inject]) 1));
val succ_in_succ = result();
@@ -117,26 +93,15 @@
goalw thy [restrict_def]
"!!f. [| restrict(f, domain(g)) = g; x: domain(g) |] ==> f`x = g`x";
by (etac subst 1);
-by (asm_full_simp_tac AC_ss 1);
+by (Asm_full_simp_tac 1);
val restrict_eq_imp_val_eq = result();
-goal thy "!!a. succ(a) = succ(b) ==> a = b";
-by (etac equalityE 1);
-by (rtac equalityI 1);
-by (fast_tac (AC_cs addSEs [mem_irrefl] addEs [mem_asym]) 1);
-by (fast_tac (AC_cs addSEs [mem_irrefl] addEs [mem_asym]) 1);
-val succ_eqD = result();
-
-goal thy "!!n. n:nat ==> 0:succ(n)";
-by (fast_tac (AC_cs addSDs [[Ord_0, nat_into_Ord] MRS Ord_linear]) 1);
-val nat_0_in_succ = result();
-
goal thy "!!f. [| domain(f)=A; f:B->C |] ==> f:A->C";
by (forward_tac [domain_of_fun] 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val domain_eq_imp_fun_type = result();
goal thy "!!R. [| R <= A * B; R ~= 0 |] ==> EX x. x:domain(R)";
-by (fast_tac (AC_cs addSEs [not_emptyE]) 1);
+by (fast_tac (!claset addSEs [not_emptyE]) 1);
val ex_in_domain = result();
--- a/src/ZF/AC/DC_lemmas.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/DC_lemmas.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,4 +1,4 @@
(*Dummy theory to document dependencies *)
-DC_lemmas = AC_Equiv + first
+DC_lemmas = AC_Equiv
--- a/src/ZF/AC/HH.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/HH.ML Fri Jan 03 15:01:55 1997 +0100
@@ -18,30 +18,30 @@
\ (let z = x - (UN b:a. HH(f,x,b)) \
\ in if(f`z:Pow(z)-{0}, f`z, {x}))";
by (resolve_tac [HH_def RS def_transrec RS trans] 1);
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
val HH_def_satisfies_eq = result();
goal thy "HH(f,x,a) : Pow(x)-{0} | HH(f,x,a)={x}";
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (simp_tac (ZF_ss addsimps [Let_def, Diff_subset RS PowI]
+by (simp_tac (!simpset addsimps [Let_def, Diff_subset RS PowI]
setloop split_tac [expand_if]) 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val HH_values = result();
goal thy "!!A. B<=A ==> X-(UN a:A. P(a)) = X-(UN a:A-B. P(a))-(UN b:B. P(b))";
-by (fast_tac (AC_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
val subset_imp_Diff_eq = result();
goal thy "!!c. [| c:a-b; b<a |] ==> c=b | b<c & c<a";
by (etac ltE 1);
by (dres_inst_tac [("x","c")] Ord_linear 1);
-by (fast_tac (AC_cs addEs [Ord_in_Ord]) 1);
-by (fast_tac (AC_cs addSIs [ltI] addIs [Ord_in_Ord]) 1);
+by (fast_tac (!claset addEs [Ord_in_Ord]) 1);
+by (fast_tac (!claset addSIs [ltI] addIs [Ord_in_Ord]) 1);
val Ord_DiffE = result();
val prems = goal thy "(!!y. y:A ==> P(y) = {x}) ==> x - (UN y:A. P(y)) = x";
-by (asm_full_simp_tac (AC_ss addsimps prems) 1);
-by (fast_tac (AC_cs addSIs [equalityI] addSDs [prem]
+by (asm_full_simp_tac (!simpset addsimps prems) 1);
+by (fast_tac (!claset addSIs [equalityI] addSDs [prem]
addSEs [RepFunE, mem_irrefl]) 1);
val Diff_UN_eq_self = result();
@@ -62,23 +62,23 @@
by (res_inst_tac [("t","%z. z-?X")] subst_context 1);
by (rtac Diff_UN_eq_self 1);
by (dtac Ord_DiffE 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addEs [ltE]) 1);
+by (fast_tac (!claset addEs [ltE]) 1);
val HH_is_x_gt_too = result();
goal thy "!!a. [| HH(f,x,a) : Pow(x)-{0}; b<a |] ==> HH(f,x,b) : Pow(x)-{0}";
by (resolve_tac [HH_values RS disjE] 1 THEN (assume_tac 1));
by (dtac HH_is_x_gt_too 1 THEN (assume_tac 1));
by (dtac subst 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [mem_irrefl]) 1);
+by (fast_tac (!claset addSEs [mem_irrefl]) 1);
val HH_subset_x_lt_too = result();
goal thy "!!a. HH(f,x,a) : Pow(x)-{0} \
\ ==> HH(f,x,a) : Pow(x - (UN b:a. HH(f,x,b)))-{0}";
by (dresolve_tac [HH_def_satisfies_eq RS subst] 1);
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (asm_full_simp_tac (AC_ss addsimps [Let_def, Diff_subset RS PowI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [Let_def, Diff_subset RS PowI]) 1);
by (dresolve_tac [expand_if RS iffD1] 1);
-by (simp_tac (ZF_ss setloop split_tac [expand_if] ) 1);
+by (simp_tac (!simpset setloop split_tac [expand_if] ) 1);
by (fast_tac (subset_cs addSEs [mem_irrefl]) 1);
val HH_subset_x_imp_subset_Diff_UN = result();
@@ -86,7 +86,7 @@
by (forw_inst_tac [("P","%y. y: Pow(x)-{0}")] subst 1 THEN (assume_tac 1));
by (dres_inst_tac [("a","w")] HH_subset_x_imp_subset_Diff_UN 1);
by (dtac subst_elem 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSIs [singleton_iff RS iffD2, equals0I]) 1);
+by (fast_tac (!claset addSIs [singleton_iff RS iffD2, equals0I]) 1);
val HH_eq_arg_lt = result();
goal thy "!!x. [| HH(f,x,v)=HH(f,x,w); HH(f,x,w): Pow(x)-{0}; \
@@ -101,7 +101,7 @@
goalw thy [lepoll_def, inj_def]
"!!i. [| HH(f, x, i) : Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}";
by (res_inst_tac [("x","lam j:i. HH(f, x, j)")] exI 1);
-by (asm_simp_tac AC_ss 1);
+by (Asm_simp_tac 1);
by (fast_tac (FOL_cs addSEs [HH_eq_imp_arg_eq, Ord_in_Ord, HH_subset_x_lt_too]
addSIs [lam_type, ballI, ltI] addIs [bexI]) 1);
val HH_subset_x_imp_lepoll = result();
@@ -113,7 +113,7 @@
val HH_Hartog_is_x = result();
goal thy "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}";
-by (fast_tac (AC_cs addSIs [Ord_Hartog, HH_Hartog_is_x, LeastI]) 1);
+by (fast_tac (!claset addSIs [Ord_Hartog, HH_Hartog_is_x, LeastI]) 1);
val HH_Least_eq_x = result();
goal thy "!!a. a:(LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) : Pow(x)-{0}";
@@ -130,8 +130,8 @@
goalw thy [inj_def]
"(lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) : \
\ inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})";
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSIs [lam_type] addDs [less_Least_subset_x]
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSIs [lam_type] addDs [less_Least_subset_x]
addSEs [HH_eq_imp_arg_eq, Ord_Least RS Ord_in_Ord]) 1);
val lam_Least_HH_inj_Pow = result();
@@ -139,65 +139,52 @@
\ ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \
\ : inj(LEAST i. HH(f,x,i)={x}, {{y}. y:x})";
by (resolve_tac [lam_Least_HH_inj_Pow RS inj_strengthen_type] 1);
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSEs [RepFun_eqI]) 1);
+by (Asm_full_simp_tac 1);
val lam_Least_HH_inj = result();
-goal thy "!!A. [| A={a}; b:A |] ==> b=a";
-by (fast_tac AC_cs 1);
-val elem_of_sing_eq = result();
-
goalw thy [surj_def]
"!!x. [| x - (UN a:A. F(a)) = 0; \
\ ALL a:A. EX z:x. F(a) = {z} |] \
\ ==> (lam a:A. F(a)) : surj(A, {{y}. y:x})";
-by (asm_full_simp_tac (AC_ss addsimps [Diff_eq_0_iff]) 1);
-by (rtac conjI 1);
-by (fast_tac (AC_cs addSIs [lam_type] addSEs [RepFun_eqI]) 1);
-by (rtac ballI 1);
-by (etac RepFunE 1);
-by (dtac subsetD 1 THEN (assume_tac 1));
-by (etac UN_E 1);
-by (dtac bspec 1 THEN (assume_tac 1));
-by (etac bexE 1);
-by (rtac bexI 1 THEN (assume_tac 2));
-by (forward_tac [elem_of_sing_eq] 1 THEN (assume_tac 1));
-by (fast_tac AC_cs 1);
+by (asm_full_simp_tac (!simpset addsimps [lam_type, Diff_eq_0_iff]) 1);
+by (Step_tac 1);
+by (set_mp_tac 1);
+by (deepen_tac (eq_cs addSIs [bexI] addSEs [equalityE]) 4 1);
val lam_surj_sing = result();
goal thy "!!x. y:Pow(x)-{0} ==> x ~= 0";
-by (fast_tac (AC_cs addSIs [equals0I, singletonI RS subst_elem]
+by (fast_tac (!claset addSIs [equals0I, singletonI RS subst_elem]
addSDs [equals0D]) 1);
val not_emptyI2 = result();
goal thy "!!f. f`(x - (UN j:i. HH(f,x,j))): Pow(x - (UN j:i. HH(f,x,j)))-{0} \
\ ==> HH(f, x, i) : Pow(x) - {0}";
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (asm_full_simp_tac (AC_ss addsimps [Let_def, Diff_subset RS PowI,
+by (asm_full_simp_tac (!simpset addsimps [Let_def, Diff_subset RS PowI,
not_emptyI2 RS if_P]) 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val f_subset_imp_HH_subset = result();
val [prem] = goal thy "(!!z. z:Pow(x)-{0} ==> f`z : Pow(z)-{0}) ==> \
\ x - (UN j: (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0";
by (excluded_middle_tac "?P : {0}" 1);
-by (fast_tac AC_cs 2);
+by (Fast_tac 2);
by (dresolve_tac [Diff_subset RS PowI RS DiffI RS prem RS
f_subset_imp_HH_subset] 1);
-by (fast_tac (AC_cs addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem)]
+by (fast_tac (!claset addSDs [HH_Least_eq_x RS sym RSN (2, subst_elem)]
addSEs [mem_irrefl]) 1);
val f_subsets_imp_UN_HH_eq_x = result();
goal thy "HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j))) | HH(f,x,i)={x}";
by (resolve_tac [HH_def_satisfies_eq RS ssubst] 1);
-by (simp_tac (ZF_ss addsimps [Let_def, Diff_subset RS PowI]
+by (simp_tac (!simpset addsimps [Let_def, Diff_subset RS PowI]
setloop split_tac [expand_if]) 1);
val HH_values2 = result();
goal thy
"!!f. HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (UN j:i. HH(f,x,j)))";
by (resolve_tac [HH_values2 RS disjE] 1 THEN (assume_tac 1));
-by (fast_tac (AC_cs addSEs [equalityE, mem_irrefl]
+by (fast_tac (!claset addSEs [equalityE, mem_irrefl]
addSDs [singleton_subsetD]) 1);
val HH_subset_imp_eq = result();
@@ -207,8 +194,8 @@
by (forward_tac [HH_subset_imp_eq] 1);
by (dtac apply_type 1);
by (resolve_tac [Diff_subset RS PowI RS DiffI] 1);
-by (fast_tac (AC_cs addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1);
-by (fast_tac (AC_cs addSEs [RepFunE] addEs [ssubst]) 1);
+by (fast_tac (!claset addSDs [HH_subset_x_imp_subset_Diff_UN RS not_emptyI2]) 1);
+by (fast_tac (!claset addSEs [RepFunE] addEs [ssubst]) 1);
val f_sing_imp_HH_sing = result();
goalw thy [bij_def]
@@ -216,7 +203,7 @@
\ f : (PROD X:Pow(x)-{0}. {{z}. z:x}) |] \
\ ==> (lam a:(LEAST i. HH(f,x,i)={x}). HH(f,x,a)) \
\ : bij(LEAST i. HH(f,x,i)={x}, {{y}. y:x})";
-by (fast_tac (AC_cs addSIs [lam_Least_HH_inj, lam_surj_sing,
+by (fast_tac (!claset addSIs [lam_Least_HH_inj, lam_surj_sing,
f_sing_imp_HH_sing]) 1);
val f_sing_lam_bij = result();
--- a/src/ZF/AC/Hartog.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/Hartog.ML Fri Jan 03 15:01:55 1997 +0100
@@ -9,7 +9,7 @@
goal thy "!!X. ALL a. Ord(a) --> a:X ==> P";
by (res_inst_tac [("X1","{y:X. Ord(y)}")] (ON_class RS revcut_rl) 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Ords_in_set";
goalw thy [lepoll_def] "!!X. [| Ord(a); a lepoll X |] ==> \
@@ -31,10 +31,10 @@
goal thy "!!X. [| Ord(a); a lepoll X |] ==> \
\ EX Y. Y<=X & (EX R. R<=X*X & ordertype(Y,R)=a)";
by (dtac Ord_lepoll_imp_ex_well_ord 1 THEN (assume_tac 1));
-by (step_tac ZF_cs 1);
+by (Step_tac 1);
by (REPEAT (ares_tac [exI, conjI] 1));
by (etac ordertype_Int 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Ord_lepoll_imp_eq_ordertype";
goal thy "!!X. ALL a. Ord(a) --> a lepoll X ==> \
@@ -44,7 +44,7 @@
by (REPEAT (eresolve_tac [allE, impE] 1));
by (assume_tac 1);
by (dtac Ord_lepoll_imp_eq_ordertype 1 THEN (assume_tac 1));
-by (fast_tac (ZF_cs addSIs [ReplaceI] addEs [sym]) 1);
+by (fast_tac (!claset addSIs [ReplaceI] addEs [sym]) 1);
qed "Ords_lepoll_set_lemma";
goal thy "!!X. ALL a. Ord(a) --> a lepoll X ==> P";
@@ -53,15 +53,15 @@
goal thy "EX a. Ord(a) & ~a lepoll X";
by (rtac swap 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (rtac Ords_lepoll_set 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "ex_Ord_not_lepoll";
goalw thy [Hartog_def] "~ Hartog(A) lepoll A";
by (resolve_tac [ex_Ord_not_lepoll RS exE] 1);
by (rtac LeastI 1);
-by (REPEAT (fast_tac ZF_cs 1));
+by (REPEAT (Fast_tac 1));
qed "HartogI";
val HartogE = HartogI RS notE;
@@ -71,14 +71,14 @@
qed "Ord_Hartog";
goalw thy [Hartog_def] "!!i. [| i < Hartog(A); ~ i lepoll A |] ==> P";
-by (fast_tac (ZF_cs addEs [less_LeastE]) 1);
+by (fast_tac (!claset addEs [less_LeastE]) 1);
qed "less_HartogE1";
goal thy "!!i. [| i < Hartog(A); i eqpoll Hartog(A) |] ==> P";
-by (fast_tac (ZF_cs addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll
+by (fast_tac (!claset addEs [less_HartogE1, eqpoll_sym RS eqpoll_imp_lepoll
RS lepoll_trans RS HartogE]) 1);
qed "less_HartogE";
goal thy "Card(Hartog(A))";
-by (fast_tac (ZF_cs addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1);
+by (fast_tac (!claset addSIs [CardI, Ord_Hartog] addEs [less_HartogE]) 1);
qed "Card_Hartog";
--- a/src/ZF/AC/OrdQuant.ML Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,104 +0,0 @@
-(* Title: ZF/AC/OrdQuant.thy
- ID: $Id$
- Authors: Krzysztof Grabczewski and L C Paulson
-
-Quantifiers and union operator for ordinals.
-*)
-
-open OrdQuant;
-
-(*** universal quantifier for ordinals ***)
-
-qed_goalw "oallI" thy [oall_def]
- "[| !!x. x<A ==> P(x) |] ==> ALL x<A. P(x)"
- (fn prems=> [ (REPEAT (ares_tac (prems @ [allI,impI]) 1)) ]);
-
-qed_goalw "ospec" thy [oall_def]
- "[| ALL x<A. P(x); x<A |] ==> P(x)"
- (fn major::prems=>
- [ (rtac (major RS spec RS mp) 1),
- (resolve_tac prems 1) ]);
-
-qed_goalw "oallE" thy [oall_def]
- "[| ALL x<A. P(x); P(x) ==> Q; ~x<A ==> Q |] ==> Q"
- (fn major::prems=>
- [ (rtac (major RS allE) 1),
- (REPEAT (eresolve_tac (prems@[asm_rl,impCE]) 1)) ]);
-
-qed_goal "rev_oallE" thy
- "[| ALL x<A. P(x); ~x<A ==> Q; P(x) ==> Q |] ==> Q"
- (fn major::prems=>
- [ (rtac (major RS oallE) 1),
- (REPEAT (eresolve_tac prems 1)) ]);
-
-(*Trival rewrite rule; (ALL x<a.P)<->P holds only if a is not 0!*)
-qed_goal "oall_simp" thy "(ALL x<a. True) <-> True"
- (fn _=> [ (REPEAT (ares_tac [TrueI,oallI,iffI] 1)) ]);
-
-(*Congruence rule for rewriting*)
-qed_goalw "oall_cong" thy [oall_def]
- "[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) |] ==> oall(a,P) <-> oall(a',P')"
- (fn prems=> [ (simp_tac (FOL_ss addsimps prems) 1) ]);
-
-
-(*** existential quantifier for ordinals ***)
-
-qed_goalw "oexI" thy [oex_def]
- "[| P(x); x<A |] ==> EX x<A. P(x)"
- (fn prems=> [ (REPEAT (ares_tac (prems @ [exI,conjI]) 1)) ]);
-
-(*Not of the general form for such rules; ~EX has become ALL~ *)
-qed_goal "oexCI" thy
- "[| ALL x<A. ~P(x) ==> P(a); a<A |] ==> EX x<A.P(x)"
- (fn prems=>
- [ (rtac classical 1),
- (REPEAT (ares_tac (prems@[oexI,oallI,notI,notE]) 1)) ]);
-
-qed_goalw "oexE" thy [oex_def]
- "[| EX x<A. P(x); !!x. [| x<A; P(x) |] ==> Q \
-\ |] ==> Q"
- (fn major::prems=>
- [ (rtac (major RS exE) 1),
- (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]);
-
-qed_goalw "oex_cong" thy [oex_def]
- "[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) \
-\ |] ==> oex(a,P) <-> oex(a',P')"
- (fn prems=> [ (simp_tac (FOL_ss addsimps prems addcongs [conj_cong]) 1) ]);
-
-
-(*** Rules for Ordinal-Indexed Unions ***)
-
-qed_goalw "OUN_I" thy [OUnion_def]
- "!!i. [| a<i; b: B(a) |] ==> b: (UN z<i. B(z))"
- (fn _=> [ fast_tac (ZF_cs addSEs [ltE]) 1 ]);
-
-qed_goalw "OUN_E" thy [OUnion_def]
- "[| b : (UN z<i. B(z)); !!a.[| b: B(a); a<i |] ==> R |] ==> R"
- (fn major::prems=>
- [ (rtac (major RS CollectE) 1),
- (rtac UN_E 1),
- (REPEAT (ares_tac (ltI::prems) 1)) ]);
-
-qed_goalw "OUN_iff" thy [oex_def]
- "b : (UN x<i. B(x)) <-> (EX x<i. b : B(x))"
- (fn _=> [ (fast_tac (FOL_cs addIs [OUN_I] addSEs [OUN_E]) 1) ]);
-
-qed_goal "OUN_cong" thy
- "[| i=j; !!x. x<j ==> C(x)=D(x) |] ==> (UN x<i.C(x)) = (UN x<j.D(x))"
- (fn prems=>
- [ rtac equality_iffI 1,
- simp_tac (ZF_ss addcongs [oex_cong] addsimps (OUN_iff::prems)) 1 ]);
-
-val OrdQuant_cs = ZF_cs
- addSIs [oallI]
- addIs [oexI, OUN_I]
- addSEs [oexE, OUN_E]
- addEs [rev_oallE];
-
-val Ord_atomize = atomize (("oall", [ospec])::ZF_conn_pairs,
- ZF_mem_pairs);
-
-val OrdQuant_ss = ZF_ss setmksimps (map mk_meta_eq o Ord_atomize o gen_all)
- addsimps [oall_simp, ltD RS beta]
- addcongs [oall_cong, oex_cong, OUN_cong];
--- a/src/ZF/AC/OrdQuant.thy Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,38 +0,0 @@
-(* Title: ZF/AC/OrdQuant.thy
- ID: $Id$
- Authors: Krzysztof Grabczewski and L C Paulson
-
-Quantifiers and union operator for ordinals.
-*)
-
-OrdQuant = Ordinal +
-
-consts
-
- (* Ordinal Quantifiers *)
- oall, oex :: [i, i => o] => o
-
- (* Ordinal Union *)
- OUnion :: [i, i => i] => i
-
-syntax
-
- "@oall" :: [idt, i, o] => o ("(3ALL _<_./ _)" 10)
- "@oex" :: [idt, i, o] => o ("(3EX _<_./ _)" 10)
- "@OUNION" :: [idt, i, i] => i ("(3UN _<_./ _)" 10)
-
-translations
-
- "ALL x<a. P" == "oall(a, %x. P)"
- "EX x<a. P" == "oex(a, %x. P)"
- "UN x<a. B" == "OUnion(a, %x. B)"
-
-defs
-
- (* Ordinal Quantifiers *)
- oall_def "oall(A, P) == ALL x. x<A --> P(x)"
- oex_def "oex(A, P) == EX x. x<A & P(x)"
-
- OUnion_def "OUnion(i,B) == {z: UN x:i. B(x). Ord(i)}"
-
-end
--- a/src/ZF/AC/Transrec2.ML Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,33 +0,0 @@
-(* Title: ZF/AC/Transrec2.ML
- ID: $Id$
- Author: Krzysztof Grabczewski
-
-Transfinite recursion introduced to handle definitions based on the three
-cases of ordinals.
-*)
-
-open Transrec2;
-
-goal thy "transrec2(0,a,b) = a";
-by (rtac (transrec2_def RS def_transrec RS trans) 1);
-by (simp_tac ZF_ss 1);
-val transrec2_0 = result();
-
-goal thy "(THE j. i=j) = i";
-by (fast_tac (ZF_cs addSIs [the_equality]) 1);
-val THE_eq = result();
-
-goal thy "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))";
-by (rtac (transrec2_def RS def_transrec RS trans) 1);
-by (simp_tac (ZF_ss addsimps [succ_not_0, THE_eq, if_P]
- setsolver K (fast_tac FOL_cs)) 1);
-val transrec2_succ = result();
-
-goal thy "!!i. Limit(i) ==> transrec2(i,a,b) = (UN j<i. transrec2(j,a,b))";
-by (rtac (transrec2_def RS def_transrec RS trans) 1);
-by (resolve_tac [if_not_P RS trans] 1 THEN
- fast_tac (OrdQuant_cs addSDs [Limit_has_0] addSEs [ltE]) 1);
-by (resolve_tac [if_not_P RS trans] 1 THEN
- fast_tac (OrdQuant_cs addSEs [succ_LimitE]) 1);
-by (simp_tac OrdQuant_ss 1);
-val transrec2_Limit = result();
--- a/src/ZF/AC/Transrec2.thy Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,23 +0,0 @@
-(* Title: ZF/AC/Transrec2.thy
- ID: $Id$
- Author: Krzysztof Grabczewski
-
-Transfinite recursion introduced to handle definitions based on the three
-cases of ordinals.
-*)
-
-Transrec2 = OrdQuant + Epsilon +
-
-consts
-
- transrec2 :: [i, i, [i,i]=>i] =>i
-
-defs
-
- transrec2_def "transrec2(alpha, a, b) ==
- transrec(alpha, %i r. if(i=0,
- a, if(EX j. i=succ(j),
- b(THE j. i=succ(j), r`(THE j. i=succ(j))),
- UN j<i. r`j)))"
-
-end
--- a/src/ZF/AC/WO1_AC.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO1_AC.ML Fri Jan 03 15:01:55 1997 +0100
@@ -32,7 +32,7 @@
(* ********************************************************************** *)
goalw thy [AC1_def, WO1_def] "!!Z. WO1 ==> AC1";
-by (fast_tac (AC_cs addSEs [ex_choice_fun]) 1);
+by (fast_tac (!claset addSEs [ex_choice_fun]) 1);
qed "WO1_AC1";
(* ********************************************************************** *)
@@ -44,19 +44,19 @@
by (eres_inst_tac [("x","Union({{C:D(B). P(C,B)}. B:A})")] allE 1);
by (etac exE 1);
by (dtac ex_choice_fun 1);
-by (fast_tac (AC_cs addEs [RepFunE, sym RS equals0D]) 1);
+by (fast_tac (!claset addEs [RepFunE, sym RS equals0D]) 1);
by (etac exE 1);
by (res_inst_tac [("x","lam x:A. f`{C:D(x). P(C,x)}")] exI 1);
-by (asm_full_simp_tac AC_ss 1);
-by (fast_tac (AC_cs addSDs [RepFunI RSN (2, apply_type)]
+by (Asm_full_simp_tac 1);
+by (fast_tac (!claset addSDs [RepFunI RSN (2, apply_type)]
addSEs [CollectD2]) 1);
val lemma1 = result();
goalw thy [WO1_def] "!!A. [| ~Finite(B); WO1 |] ==> |B| + |B| eqpoll B";
by (rtac eqpoll_trans 1);
-by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 2);
+by (fast_tac (!claset addSEs [well_ord_cardinal_eqpoll]) 2);
by (resolve_tac [eqpoll_sym RS eqpoll_trans] 1);
-by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 1);
+by (fast_tac (!claset addSEs [well_ord_cardinal_eqpoll]) 1);
by (resolve_tac [cadd_def RS def_imp_eq RS subst] 1);
by (resolve_tac [Card_cardinal RSN (2, Inf_Card_is_InfCard) RS
InfCard_cdouble_eq RS ssubst] 1);
@@ -65,23 +65,23 @@
by (etac notE 1);
by (resolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_Finite] 1
THEN (assume_tac 2));
-by (fast_tac (AC_cs addSEs [well_ord_cardinal_eqpoll]) 1);
+by (fast_tac (!claset addSEs [well_ord_cardinal_eqpoll]) 1);
val lemma2_1 = result();
goal thy "!!f. f : bij(D+D, B) ==> {{f`Inl(i), f`Inr(i)}. i:D} : Pow(Pow(B))";
-by (fast_tac (AC_cs addSIs [InlI, InrI]
+by (fast_tac (!claset addSIs [InlI, InrI]
addSEs [RepFunE, bij_is_fun RS apply_type]) 1);
val lemma2_2 = result();
goal thy "!!f. [| f:inj(A,B); f`a = f`b; a:A; b:A |] ==> a=b";
by (rtac inj_equality 1);
-by (TRYALL (fast_tac (AC_cs addSEs [inj_is_fun RS apply_Pair] addEs [subst])));
+by (TRYALL (fast_tac (!claset addSEs [inj_is_fun RS apply_Pair] addEs [subst])));
val lemma = result();
goalw thy AC_aux_defs
"!!f. f : bij(D+D, B) ==> \
\ pairwise_disjoint({{f`Inl(i), f`Inr(i)}. i:D})";
-by (fast_tac (AC_cs addSEs [RepFunE, not_emptyE]
+by (fast_tac (!claset addSEs [RepFunE, not_emptyE]
addDs [bij_is_inj RS lemma, Inl_iff RS iffD1,
Inr_iff RS iffD1, Inl_Inr_iff RS iffD1 RS FalseE,
Inr_Inl_iff RS iffD1 RS FalseE]
@@ -92,7 +92,7 @@
"[| f : bij(D+D, B); 1 le n |] ==> \
\ sets_of_size_between({{f`Inl(i), f`Inr(i)}. i:D}, 2, succ(n))";
by (rewtac succ_def);
-by (fast_tac (AC_cs addSIs [cons_lepoll_cong, minor, lepoll_refl, InlI, InrI]
+by (fast_tac (!claset addSIs [cons_lepoll_cong, minor, lepoll_refl, InlI, InrI]
addIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
le_imp_subset RS subset_imp_lepoll]
addDs [major RS bij_is_inj RS lemma, Inl_Inr_iff RS iffD1 RS FalseE]
@@ -101,7 +101,7 @@
goalw thy [bij_def, surj_def]
"!!f. f : bij(D+D, B) ==> Union({{f`Inl(i), f`Inr(i)}. i:D})=B";
-by (fast_tac (AC_cs addSEs [inj_is_fun RS apply_type, CollectE, sumE]
+by (fast_tac (!claset addSEs [inj_is_fun RS apply_type, CollectE, sumE]
addSIs [InlI, InrI, equalityI]) 1);
val lemma2_5 = result();
@@ -116,5 +116,5 @@
val lemma2 = result();
goalw thy AC_defs "!!n. [| WO1; 1 le n |] ==> AC10(n)";
-by (fast_tac (AC_cs addSIs [lemma1] addSEs [lemma2]) 1);
+by (fast_tac (!claset addSIs [lemma1] addSEs [lemma2]) 1);
qed "WO1_AC10";
--- a/src/ZF/AC/WO1_WO6.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO1_WO6.ML Fri Jan 03 15:01:55 1997 +0100
@@ -5,43 +5,43 @@
Proofs needed to state that formulations WO1,...,WO6 are all equivalent.
All but one WO6 ==> WO1 (placed in separate file WO6_WO1.ML)
- Every proofs (exept one) presented in this file are referred as clear
+ Every proof (exept one) presented in this file are referred as "clear"
by Rubin & Rubin (page 2).
- They refer reader to a book by G"odel to see the proof WO1 ==> WO2.
+ They refer reader to a book by Gödel to see the proof WO1 ==> WO2.
Fortunately order types made this proof also very easy.
*)
(* ********************************************************************** *)
goalw thy WO_defs "!!Z. WO2 ==> WO3";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "WO2_WO3";
(* ********************************************************************** *)
goalw thy (eqpoll_def::WO_defs) "!!Z. WO3 ==> WO1";
-by (fast_tac (ZF_cs addSEs [bij_is_inj RS well_ord_rvimage,
- well_ord_Memrel RS well_ord_subset]) 1);
+by (fast_tac (!claset addSEs [bij_is_inj RS well_ord_rvimage,
+ well_ord_Memrel RS well_ord_subset]) 1);
qed "WO3_WO1";
(* ********************************************************************** *)
goalw thy (eqpoll_def::WO_defs) "!!Z. WO1 ==> WO2";
-by (fast_tac (ZF_cs addSIs [Ord_ordertype, ordermap_bij]) 1);
+by (fast_tac (!claset addSIs [Ord_ordertype, ordermap_bij]) 1);
qed "WO1_WO2";
(* ********************************************************************** *)
goal thy "!!f. f: A->B ==> (lam x:A. {f`x}): A -> {{b}. b:B}";
-by (fast_tac (ZF_cs addSIs [lam_type, apply_type]) 1);
+by (fast_tac (!claset addSIs [lam_type, apply_type]) 1);
val lam_sets = result();
goalw thy [surj_def] "!!f. f:surj(A,B) ==> (UN a:A. {f`a}) = B";
-by (fast_tac (ZF_cs addSIs [equalityI] addSEs [apply_type]) 1);
+by (fast_tac (!claset addSIs [equalityI] addSEs [apply_type]) 1);
val surj_imp_eq_ = result();
goal thy "!!f. [| f:surj(A,B); Ord(A) |] ==> (UN a<A. {f`a}) = B";
-by (fast_tac (AC_cs addSDs [surj_imp_eq_]
+by (fast_tac (!claset addSDs [surj_imp_eq_]
addSIs [equalityI, ltI] addSEs [ltE]) 1);
val surj_imp_eq = result();
@@ -54,7 +54,7 @@
by (rtac conjI 1);
by (eresolve_tac [ordermap_bij RS bij_converse_bij RS bij_is_fun RS
lam_sets RS domain_of_fun] 1);
-by (asm_simp_tac (AC_ss addsimps [singleton_eqpoll_1 RS eqpoll_imp_lepoll,
+by (asm_simp_tac (!simpset addsimps [singleton_eqpoll_1 RS eqpoll_imp_lepoll,
Ord_ordertype RSN (2, ordermap_bij RS bij_converse_bij RS
bij_is_surj RS surj_imp_eq)]) 1);
qed "WO1_WO4";
@@ -62,18 +62,20 @@
(* ********************************************************************** *)
goalw thy WO_defs "!!Z. [| m:nat; n:nat; m le n; WO4(m) |] ==> WO4(n)";
-by (fast_tac (AC_cs addIs [nat_le_imp_lepoll RSN (2, lepoll_trans)]) 1);
+by (fast_tac (!claset addIs [nat_le_imp_lepoll RSN (2, lepoll_trans)]) 1);
qed "WO4_mono";
(* ********************************************************************** *)
goalw thy WO_defs "!!Z. [| m:nat; 1 le m; WO4(m) |] ==> WO5";
+ (*ZF_cs is essential: default claset's too slow*)
by (fast_tac ZF_cs 1);
qed "WO4_WO5";
(* ********************************************************************** *)
goalw thy WO_defs "!!Z. WO5 ==> WO6";
+ (*ZF_cs is essential: default claset's too slow*)
by (fast_tac ZF_cs 1);
qed "WO5_WO6";
--- a/src/ZF/AC/WO1_WO7.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO1_WO7.ML Fri Jan 03 15:01:55 1997 +0100
@@ -12,7 +12,7 @@
goalw thy [WO7_def] "WO7 <-> (ALL X. ~Finite(X) --> \
\ (EX R. well_ord(X,R) & ~well_ord(X,converse(R))))";
-by (fast_tac (ZF_cs addSEs [Finite_well_ord_converse]) 1);
+by (fast_tac (!claset addSEs [Finite_well_ord_converse]) 1);
val WO7_iff_LEMMA = result();
(* ********************************************************************** *)
@@ -24,9 +24,9 @@
by (rtac allI 1);
by (etac allE 1);
by (excluded_middle_tac "Finite(A)" 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (rewrite_goals_tac [Finite_def, eqpoll_def]);
-by (fast_tac (ZF_cs addSIs [[bij_is_inj, nat_implies_well_ord] MRS
+by (fast_tac (!claset addSIs [[bij_is_inj, nat_implies_well_ord] MRS
well_ord_rvimage]) 1);
val LEMMA_imp_WO1 = result();
@@ -50,16 +50,16 @@
by (rtac notI 1);
by (eres_inst_tac [("x","nat")] allE 1);
by (etac disjE 1);
-by (fast_tac (ZF_cs addSDs [nat_0I RSN (2,equals0D)]) 1);
+by (fast_tac (!claset addSDs [nat_0I RSN (2,equals0D)]) 1);
by (etac bexE 1);
by (eres_inst_tac [("x","succ(x)")] allE 1);
-by (fast_tac (ZF_cs addSIs [nat_succI, converseI, MemrelI,
+by (fast_tac (!claset addSIs [nat_succI, converseI, MemrelI,
nat_succI RSN (2, subsetD)]) 1);
val converse_Memrel_not_wf_on = result();
goalw thy [well_ord_def]
"!!a. [| Ord(a); ~Finite(a) |] ==> ~well_ord(a,converse(Memrel(a)))";
-by (fast_tac (ZF_cs addSDs [converse_Memrel_not_wf_on]) 1);
+by (fast_tac (!claset addSDs [converse_Memrel_not_wf_on]) 1);
val converse_Memrel_not_well_ord = result();
goal thy "!!A. [| well_ord(A,r); well_ord(A,converse(r)) |] \
--- a/src/ZF/AC/WO1_WO8.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO1_WO8.ML Fri Jan 03 15:01:55 1997 +0100
@@ -10,7 +10,7 @@
(* ********************************************************************** *)
goalw thy WO_defs "!!Z. WO1 ==> WO8";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "WO1_WO8";
(* ********************************************************************** *)
@@ -21,10 +21,10 @@
by (rtac allI 1);
by (eres_inst_tac [("x","{{x}. x:A}")] allE 1);
by (etac impE 1);
-by (fast_tac (AC_cs addSEs [lam_sing_bij RS bij_is_inj RS
+by (fast_tac (!claset addSEs [lam_sing_bij RS bij_is_inj RS
well_ord_rvimage]) 2);
by (res_inst_tac [("x","lam a:{{x}. x:A}. THE x. a={x}")] exI 1);
-by (fast_tac (ZF_cs addSEs [singleton_eq_iff RS iffD1 RS sym]
+by (fast_tac (!claset addSEs [singleton_eq_iff RS iffD1 RS sym]
addSIs [lam_type]
addIs [the_equality RS ssubst]) 1);
qed "WO8_WO1";
--- a/src/ZF/AC/WO2_AC16.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO2_AC16.ML Fri Jan 03 15:01:55 1997 +0100
@@ -42,18 +42,17 @@
by (etac lemma3_1 1 THEN (REPEAT (assume_tac 1)));
val lemma3 = result();
-goal thy "!!a. [| ALL y:x. y < a --> F(y) <= X & \
+goal thy "!!a. [| ALL y<x. F(y) <= X & \
\ (ALL x<a. x < y | (EX Y:F(y). fa(x) <= Y) --> \
\ (EX! Y. Y : F(y) & fa(x) <= Y)); x < a |] \
\ ==> ALL y<x. ALL z<a. z < y | (EX Y:F(y). fa(z) <= Y) --> \
\ (EX! Y. Y : F(y) & fa(z) <= Y)";
by (REPEAT (resolve_tac [oallI, impI] 1));
-by (dresolve_tac [ltD RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (eresolve_tac [lt_trans RSN (2, impE)] 1 THEN (REPEAT (assume_tac 1)));
+by (dresolve_tac [ospec] 1 THEN (assume_tac 1));
by (fast_tac (FOL_cs addSEs [oallE]) 1);
val lemma4 = result();
-goal thy "!!a. [| ALL y:x. y < a --> F(y) <= X & \
+goal thy "!!a. [| ALL y<x. F(y) <= X & \
\ (ALL x<a. x < y | (EX Y:F(y). fa(x) <= Y) --> \
\ (EX! Y. Y : F(y) & fa(x) <= Y)); \
\ x < a; Limit(x); ALL i j. i le j --> F(i) <= F(j) |] \
@@ -63,9 +62,8 @@
by (rtac conjI 1);
by (rtac subsetI 1);
by (etac OUN_E 1);
-by (dresolve_tac [ltD RSN (2, bspec)] 1 THEN (assume_tac 1));
-by (eresolve_tac [lt_trans RSN (2, impE)] 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac AC_cs 1);
+by (dresolve_tac [ospec] 1 THEN (assume_tac 1));
+by (Fast_tac 1);
by (dtac lemma4 1 THEN (assume_tac 1));
by (rtac oallI 1);
by (rtac impI 1);
@@ -119,9 +117,9 @@
by (dresolve_tac [nat_le_infinite_Ord RS le_imp_lepoll] 1
THEN (assume_tac 1));
by (rewtac Finite_def);
-by (fast_tac (AC_cs addSEs [eqpoll_sym RS eqpoll_trans]) 2);
+by (fast_tac (!claset addSEs [eqpoll_sym RS eqpoll_trans]) 2);
by (rtac lepoll_trans 1 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_subset RS
+by (fast_tac (!claset addSEs [Ord_nat RSN (2, ltI) RS leI RS le_imp_subset RS
subset_imp_lepoll RSN (2, eqpoll_imp_lepoll RS lepoll_trans)]) 1);
val Finite_lesspoll_infinite_Ord = result();
@@ -132,25 +130,25 @@
goal thy "!!n. n:nat ==> ALL X. X eqpoll n --> (ALL x:X. Finite(x)) \
\ --> Finite(Union(X))";
by (etac nat_induct 1);
-by (fast_tac (AC_cs addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]
- addSIs [nat_0I RS nat_into_Finite] addIs [Union_0 RS ssubst]) 1);
+by (fast_tac (!claset addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]
+ addSIs [nat_0I RS nat_into_Finite] addss (!simpset)) 1);
by (REPEAT (resolve_tac [allI, impI] 1));
by (resolve_tac [eqpoll_succ_imp_not_empty RS not_emptyE] 1 THEN (assume_tac 1));
by (res_inst_tac [("P","%z. Finite(z)")] (Union_eq_Un_Diff RS ssubst) 1
THEN (assume_tac 1));
by (rtac Finite_Un 1);
-by (fast_tac AC_cs 2);
-by (fast_tac (AC_cs addSIs [Diff_sing_eqpoll]) 1);
+by (Fast_tac 2);
+by (fast_tac (!claset addSIs [Diff_sing_eqpoll]) 1);
val Finite_Union_lemma = result();
goal thy "!!X. [| ALL x:X. Finite(x); Finite(X) |] ==> Finite(Union(X))";
by (eresolve_tac [Finite_def RS def_imp_iff RS iffD1 RS bexE] 1);
by (dtac Finite_Union_lemma 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
val Finite_Union = result();
goalw thy [Finite_def] "!!x. [| x lepoll n; n:nat |] ==> Finite(x)";
-by (fast_tac (AC_cs
+by (fast_tac (!claset
addEs [nat_into_Ord RSN (2, lepoll_imp_ex_le_eqpoll) RS exE,
Ord_nat RSN (2, ltI) RSN (2, lt_trans1) RS ltD]) 1);
val lepoll_nat_num_imp_Finite = result();
@@ -161,7 +159,7 @@
by (excluded_middle_tac "Finite(X)" 1);
by (resolve_tac [Card_is_Ord RSN (3, Finite_lesspoll_infinite_Ord)] 2
THEN (REPEAT (assume_tac 3)));
-by (fast_tac (AC_cs addSEs [lepoll_nat_num_imp_Finite]
+by (fast_tac (!claset addSEs [lepoll_nat_num_imp_Finite]
addSIs [Finite_Union]) 2);
by (dresolve_tac [lt_Ord RSN (2, lepoll_imp_ex_le_eqpoll)] 1 THEN (assume_tac 1));
by (REPEAT (eresolve_tac [exE, conjE] 1));
@@ -176,7 +174,7 @@
by (eresolve_tac [lt_trans1 RSN (2, lt_Card_imp_lesspoll)] 1
THEN REPEAT (assume_tac 1));
by (rtac UN_lepoll 1
- THEN (TRYALL (fast_tac (AC_cs addSEs [lt_Ord]))));
+ THEN (TRYALL (fast_tac (!claset addSEs [lt_Ord]))));
val Union_lesspoll = result();
(* ********************************************************************** *)
@@ -188,36 +186,36 @@
val Un_sing_eq_cons = result();
goal thy "!!A. A lepoll B ==> A Un {a} lepoll succ(B)";
-by (asm_simp_tac (AC_ss addsimps [Un_sing_eq_cons, succ_def]) 1);
+by (asm_simp_tac (!simpset addsimps [Un_sing_eq_cons, succ_def]) 1);
by (eresolve_tac [mem_not_refl RSN (2, cons_lepoll_cong)] 1);
val Un_lepoll_succ = result();
goal thy "!!a. Ord(a) ==> F(a) - (UN b<succ(a). F(b)) = 0";
-by (fast_tac (AC_cs addSIs [OUN_I, le_refl] addIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [OUN_I, le_refl] addIs [equalityI]) 1);
val Diff_UN_succ_empty = result();
goal thy "!!a. Ord(a) ==> F(a) Un X - (UN b<succ(a). F(b)) <= X";
-by (fast_tac (AC_cs addSIs [OUN_I, le_refl]) 1);
+by (fast_tac (!claset addSIs [OUN_I, le_refl]) 1);
val Diff_UN_succ_subset = result();
goal thy "!!x. Ord(x) ==> \
\ recfunAC16(f, g, x, a) - (UN i<x. recfunAC16(f, g, i, a)) lepoll 1";
by (etac Ord_cases 1);
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_0,
+by (asm_simp_tac (!simpset addsimps [recfunAC16_0,
empty_subsetI RS subset_imp_lepoll]) 1);
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_Limit,
+by (asm_simp_tac (!simpset addsimps [recfunAC16_Limit,
Diff_cancel, empty_subsetI RS subset_imp_lepoll]) 2);
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_succ]) 1);
+by (asm_simp_tac (!simpset addsimps [recfunAC16_succ]) 1);
by (resolve_tac [conjI RS (expand_if RS iffD2)] 1);
-by (fast_tac (AC_cs addSIs [empty_subsetI RS subset_imp_lepoll]
+by (fast_tac (!claset addSIs [empty_subsetI RS subset_imp_lepoll]
addSEs [Diff_UN_succ_empty RS ssubst]) 1);
-by (fast_tac (AC_cs addSEs [Diff_UN_succ_subset RS subset_imp_lepoll RS
+by (fast_tac (!claset addSEs [Diff_UN_succ_subset RS subset_imp_lepoll RS
(singleton_eqpoll_1 RS eqpoll_imp_lepoll RSN (2, lepoll_trans))]) 1);
val recfunAC16_Diff_lepoll_1 = result();
goal thy "!!z. [| z : F(x); Ord(x) |] \
\ ==> z:F(LEAST i. z:F(i)) - (UN j<(LEAST i. z:F(i)). F(j))";
-by (fast_tac (AC_cs addEs [less_LeastE] addSEs [OUN_E, LeastI]) 1);
+by (fast_tac (!claset addEs [less_LeastE] addSEs [OUN_E, LeastI]) 1);
val in_Least_Diff = result();
goal thy "!!w. [| (LEAST i. w:F(i)) = (LEAST i. z:F(i)); \
@@ -234,7 +232,7 @@
val Least_eq_imp_ex = result();
goal thy "!!A. [| A lepoll 1; a:A; b:A |] ==> a=b";
-by (fast_tac (AC_cs addSDs [lepoll_1_is_sing]) 1);
+by (fast_tac (!claset addSDs [lepoll_1_is_sing]) 1);
val two_in_lepoll_1 = result();
goal thy "!!a. [| ALL i<a. F(i)-(UN j<i. F(j)) lepoll 1; Limit(a) |] \
@@ -250,23 +248,23 @@
by (etac lt_Ord2 1);
by (rtac ballI 1);
by (rtac ballI 1);
-by (asm_simp_tac AC_ss 1);
+by (Asm_simp_tac 1);
by (rtac impI 1);
by (dtac Least_eq_imp_ex 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac (AC_cs addSEs [two_in_lepoll_1]) 1);
+by (fast_tac (!claset addSEs [two_in_lepoll_1]) 1);
val UN_lepoll_index = result();
goal thy "!!y. Ord(y) ==> recfunAC16(f, fa, y, a) lepoll y";
by (etac trans_induct 1);
by (etac Ord_cases 1);
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_0, lepoll_refl]) 1);
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_succ]) 1);
-by (fast_tac (AC_cs addIs [conjI RS (expand_if RS iffD2)]
+by (asm_simp_tac (!simpset addsimps [recfunAC16_0, lepoll_refl]) 1);
+by (asm_simp_tac (!simpset addsimps [recfunAC16_succ]) 1);
+by (fast_tac (!claset addIs [conjI RS (expand_if RS iffD2)]
addSDs [succI1 RSN (2, bspec)]
addSEs [subset_succI RS subset_imp_lepoll RSN (2, lepoll_trans),
Un_lepoll_succ]) 1);
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_Limit]) 1);
-by (fast_tac (AC_cs addSEs [lt_Ord RS recfunAC16_Diff_lepoll_1]
+by (asm_simp_tac (!simpset addsimps [recfunAC16_Limit]) 1);
+by (fast_tac (!claset addSEs [lt_Ord RS recfunAC16_Diff_lepoll_1]
addSIs [UN_lepoll_index]) 1);
val recfunAC16_lepoll_index = result();
@@ -278,7 +276,7 @@
by (eresolve_tac [lt_Ord RS recfunAC16_lepoll_index] 3);
by (eresolve_tac [[bij_is_inj, Card_is_Ord RS well_ord_Memrel] MRS
well_ord_rvimage] 2 THEN (assume_tac 2));
-by (fast_tac (AC_cs addSEs [eqpoll_imp_lepoll]) 1);
+by (fast_tac (!claset addSEs [eqpoll_imp_lepoll]) 1);
val Union_recfunAC16_lesspoll = result();
goal thy
@@ -308,11 +306,11 @@
\ fa : bij(a, {x: Pow(A) . x eqpoll k}); i<a; k:nat; m:nat |] \
\ ==> fa ` i Un x : {x: Pow(A) . x eqpoll k #+ m}";
by (rtac CollectI 1);
-by (fast_tac (AC_cs addSIs [PowD RS (PowD RSN (2, Un_least RS PowI))]
+by (fast_tac (!claset addSIs [PowD RS (PowD RSN (2, Un_least RS PowI))]
addSEs [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)]) 1);
by (rtac disj_Un_eqpoll_nat_sum 1
THEN (TRYALL assume_tac));
-by (fast_tac (AC_cs addSIs [equals0I]) 1);
+by (fast_tac (!claset addSIs [equals0I]) 1);
by (eresolve_tac [ltD RSN (2, bij_is_fun RS apply_type RS CollectE)] 1
THEN (REPEAT (assume_tac 1)));
val Un_in_Collect = result();
@@ -321,18 +319,17 @@
(* Lemmas simplifying assumptions *)
(* ********************************************************************** *)
-goal thy "!!j. [| ALL y:succ(j). y<a --> F(y)<=X & (ALL x<a. x<y | P(x,y) \
+goal thy "!!j. [| ALL y<succ(j). F(y)<=X & (ALL x<a. x<y | P(x,y) \
\ --> Q(x,y)); succ(j)<a |] \
\ ==> F(j)<=X & (ALL x<a. x<j | P(x,j) --> Q(x,j))";
-by (dresolve_tac [succI1 RSN (2, bspec)] 1);
-by (etac impE 1);
-by (resolve_tac [lt_Ord RS (succI1 RS ltI RS lt_Ord RS le_refl) RS lt_trans] 1
+by (dresolve_tac [ospec] 1);
+by (resolve_tac [lt_Ord RS (succI1 RS ltI RS lt_Ord RS le_refl)] 1
THEN (REPEAT (assume_tac 1)));
val lemma6 = result();
goal thy "!!j. [| F(j)<=X; (ALL x<a. x<j | P(x,j) --> Q(x,j)); succ(j)<a |] \
\ ==> P(j,j) --> F(j) <= X & (ALL x<a. x le j | P(x,j) --> Q(x,j))";
-by (fast_tac (AC_cs addSEs [leE]) 1);
+by (fast_tac (!claset addSEs [leE]) 1);
val lemma7 = result();
(* ********************************************************************** *)
@@ -348,15 +345,15 @@
((eqpoll_sym RS eqpoll_imp_lepoll) RSN (2, lepoll_trans)) RS
lepoll_imp_eqpoll_subset RS exE] 1
THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSEs [eqpoll_sym]) 1);
+by (fast_tac (!claset addSEs [eqpoll_sym]) 1);
val ex_subset_eqpoll = result();
goal thy "!!A. [| A <= B Un C; A Int C = 0 |] ==> A <= B";
-by (fast_tac (AC_cs addDs [equals0D]) 1);
+by (fast_tac (!claset addDs [equals0D]) 1);
val subset_Un_disjoint = result();
goal thy "!!F. [| X:Pow(A - Union(B) -C); T:B; F<=T |] ==> F Int X = 0";
-by (fast_tac (AC_cs addSIs [equals0I]) 1);
+by (fast_tac (!claset addSIs [equals0I]) 1);
val Int_empty = result();
(* ********************************************************************** *)
@@ -369,7 +366,7 @@
goal thy "!!A. m:nat ==> ALL A B. A <= B & m lepoll A & B lepoll m --> A=B";
by (etac nat_induct 1);
-by (fast_tac (AC_cs addSDs [lepoll_0_is_0] addSIs [equalityI]) 1);
+by (fast_tac (!claset addSDs [lepoll_0_is_0] addSIs [equalityI]) 1);
by (REPEAT (resolve_tac [allI, impI] 1));
by (REPEAT (etac conjE 1));
by (resolve_tac [succ_lepoll_imp_not_empty RS not_emptyE] 1
@@ -379,8 +376,8 @@
by (forward_tac [lepoll_Diff_sing] 1);
by (REPEAT (eresolve_tac [allE, impE] 1));
by (rtac conjI 1);
-by (fast_tac AC_cs 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 2);
+by (Fast_tac 1);
by (etac Diffs_eq_imp_eq 1
THEN REPEAT (assume_tac 1));
val subset_imp_eq_lemma = result();
@@ -393,12 +390,12 @@
\ y<a |] ==> b=y";
by (dtac subset_imp_eq 1);
by (etac nat_succI 3);
-by (fast_tac (AC_cs addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS
+by (fast_tac (!claset addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS
CollectE, eqpoll_sym RS eqpoll_imp_lepoll]) 1);
-by (fast_tac (AC_cs addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS
+by (fast_tac (!claset addSEs [bij_is_fun RS (ltD RSN (2, apply_type)) RS
CollectE, eqpoll_imp_lepoll]) 1);
by (rewrite_goals_tac [bij_def, inj_def]);
-by (fast_tac (AC_cs addSDs [ltD]) 1);
+by (fast_tac (!claset addSDs [ltD]) 1);
val bij_imp_arg_eq = result();
goal thy
@@ -423,7 +420,7 @@
by (etac nat_succI 1);
by (assume_tac 1);
by (rtac conjI 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (REPEAT (resolve_tac [ballI, impI, oallI, notI] 1));
by (dresolve_tac [Int_empty RSN (2, subset_Un_disjoint)] 1
THEN REPEAT (assume_tac 1));
@@ -456,7 +453,7 @@
val ex_next_Ord = result();
goal thy "!!Z. [| EX! Y. Y:Z & P(Y); ~P(W) |] ==> EX! Y. Y:Z Un {W} & P(Y)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val ex1_in_Un_sing = result();
(* ********************************************************************** *)
@@ -470,7 +467,7 @@
\ (ALL x<a. x le j | (EX xa:F(j) Un {L}. P(x, xa)) --> \
\ (EX! Y. Y:F(j) Un {L} & P(x, Y)))";
by (rtac conjI 1);
-by (fast_tac (AC_cs addSIs [singleton_subsetI]) 1);
+by (fast_tac (!claset addSIs [singleton_subsetI]) 1);
by (rtac oallI 1);
by (etac oallE 1 THEN (contr_tac 2));
by (rtac impI 1);
@@ -479,12 +476,12 @@
by (eresolve_tac [disjI1 RSN (2, impE)] 1 THEN (assume_tac 1));
by (rtac ex1E 1 THEN (assume_tac 1));
by (etac ex1_in_Un_sing 1);
-by (fast_tac AC_cs 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
+by (Fast_tac 1);
by (etac bexE 1);
by (etac UnE 1);
-by (fast_tac (AC_cs addSEs [ex1_in_Un_sing]) 1);
-by (fast_tac AC_cs 1);
+by (fast_tac (!claset addSEs [ex1_in_Un_sing]) 1);
+by (Fast_tac 1);
val lemma8 = result();
(* ********************************************************************** *)
@@ -499,21 +496,19 @@
\ ==> recfunAC16(f, fa, b, a) <= {X: Pow(A) . X eqpoll succ(k #+ m)} & \
\ (ALL x<a. x < b | (EX Y:recfunAC16(f, fa, b, a). fa ` x <= Y) --> \
\ (EX! Y. Y:recfunAC16(f, fa, b, a) & fa ` x <= Y))";
-by (rtac impE 1 THEN (REPEAT (assume_tac 2)));
-by (eresolve_tac [lt_Ord RS trans_induct] 1);
-by (rtac impI 1);
+by (etac lt_induct 1);
+by (forward_tac [lt_Ord] 1);
by (etac Ord_cases 1);
(* case 0 *)
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_0]) 1);
-by (fast_tac (AC_cs addSEs [ltE]) 1);
+by (asm_simp_tac (!simpset addsimps [recfunAC16_0]) 1);
(* case Limit *)
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_Limit]) 2);
-by (etac lemma5 2 THEN (REPEAT (assume_tac 2)));
+by (asm_simp_tac (!simpset addsimps [recfunAC16_Limit]) 2);
+by (rtac lemma5 2 THEN (REPEAT (assume_tac 2)));
by (fast_tac (FOL_cs addSEs [recfunAC16_mono]) 2);
(* case succ *)
by (hyp_subst_tac 1);
by (eresolve_tac [lemma6 RS conjE] 1 THEN (assume_tac 1));
-by (asm_simp_tac (AC_ss addsimps [recfunAC16_succ]) 1);
+by (asm_simp_tac (!simpset addsimps [recfunAC16_succ]) 1);
by (resolve_tac [conjI RS (expand_if RS iffD2)] 1);
by (etac lemma7 1 THEN (REPEAT (assume_tac 1)));
by (rtac impI 1);
@@ -523,6 +518,7 @@
by (resolve_tac [bij_is_fun RS apply_type] 1 THEN (assume_tac 1));
by (eresolve_tac [Least_le RS lt_trans2 RS ltD] 1
THEN REPEAT (ares_tac [lt_Ord, succ_leI] 1));
+(*VERY SLOW! 24 secs??*)
by (eresolve_tac [lt_Ord RSN (2, LeastI)] 1 THEN (assume_tac 1));
val main_induct = result();
@@ -541,7 +537,7 @@
by (rtac subsetI 1);
by (etac OUN_E 1);
by (dtac prem1 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (rtac ballI 1);
by (etac imageE 1);
by (dresolve_tac [prem3 RS Limit_is_Ord RSN (2, ltI) RS
@@ -550,15 +546,15 @@
by (etac conjE 1);
by (dresolve_tac [leI RS succ_leE RSN (2, ospec)] 1 THEN (assume_tac 1));
by (etac impE 1);
-by (fast_tac (AC_cs addSEs [leI RS succ_leE RS lt_Ord RS le_refl]) 1);
+by (fast_tac (!claset addSEs [leI RS succ_leE RS lt_Ord RS le_refl]) 1);
by (dresolve_tac [prem2 RSN (2, apply_equality)] 1);
by (REPEAT (eresolve_tac [conjE, ex1E] 1));
by (rtac ex1I 1);
-by (fast_tac (AC_cs addSIs [OUN_I]) 1);
+by (fast_tac (!claset addSIs [OUN_I]) 1);
by (REPEAT (eresolve_tac [conjE, OUN_E] 1));
by (eresolve_tac [lt_Ord RSN (2, lt_Ord RS Ord_linear_le)] 1 THEN (assume_tac 1));
by (dresolve_tac [prem4 RS subsetD] 2 THEN (assume_tac 2));
-by (fast_tac FOL_cs 2);
+by (Fast_tac 2);
by (forward_tac [prem1] 1);
by (forward_tac [succ_leE] 1);
by (dresolve_tac [prem4 RS subsetD] 1 THEN (assume_tac 1));
@@ -566,7 +562,7 @@
by (dresolve_tac [lt_trans RSN (2, ospec)] 1 THEN (TRYALL assume_tac));
by (dresolve_tac [disjI1 RSN (2, mp)] 1 THEN (assume_tac 1));
by (etac ex1_two_eq 1);
-by (REPEAT (fast_tac FOL_cs 1));
+by (REPEAT (Fast_tac 1));
val lemma_simp_induct = result();
(* ********************************************************************** *)
--- a/src/ZF/AC/WO6_WO1.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO6_WO1.ML Fri Jan 03 15:01:55 1997 +0100
@@ -16,12 +16,12 @@
by (res_inst_tac [("i","k"),("j","i")] Ord_linear2 1);
by (dtac odiff_lt_mono2 4 THEN assume_tac 4);
by (asm_full_simp_tac
- (ZF_ss addsimps [oadd_odiff_inverse, odiff_oadd_inverse]) 4);
-by (safe_tac (ZF_cs addSEs [lt_Ord]));
+ (!simpset addsimps [oadd_odiff_inverse, odiff_oadd_inverse]) 4);
+by (safe_tac (!claset addSEs [lt_Ord]));
val lt_oadd_odiff_disj = result();
(*The corresponding elimination rule*)
-val lt_oadd_odiff_cases = rule_by_tactic (safe_tac ZF_cs)
+val lt_oadd_odiff_cases = rule_by_tactic (safe_tac (!claset))
(lt_oadd_odiff_disj RS disjE);
(* ********************************************************************** *)
@@ -33,25 +33,25 @@
(* ********************************************************************** *)
goalw thy [uu_def] "domain(uu(f,b,g,d)) <= f`b";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val domain_uu_subset = result();
goal thy "!! a. ALL b<a. f`b lepoll m ==> \
\ ALL b<a. ALL g<a. ALL d<a. domain(uu(f,b,g,d)) lepoll m";
-by (fast_tac (AC_cs addSEs
+by (fast_tac (!claset addSEs
[domain_uu_subset RS subset_imp_lepoll RS lepoll_trans]) 1);
val quant_domain_uu_lepoll_m = result();
goalw thy [uu_def] "uu(f,b,g,d) <= f`b * f`g";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val uu_subset1 = result();
goalw thy [uu_def] "uu(f,b,g,d) <= f`d";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val uu_subset2 = result();
goal thy "!! a. [| ALL b<a. f`b lepoll m; d<a |] ==> uu(f,b,g,d) lepoll m";
-by (fast_tac (AC_cs
+by (fast_tac (!claset
addSEs [uu_subset2 RS subset_imp_lepoll RS lepoll_trans]) 1);
val uu_lepoll_m = result();
@@ -64,15 +64,15 @@
\ u(f,b,g,d) lesspoll m)) | \
\ (EX b<a. f`b ~= 0 & (ALL g<a. ALL d<a. u(f,b,g,d) ~= 0 --> \
\ u(f,b,g,d) eqpoll m))";
-by (asm_simp_tac OrdQuant_ss 1);
-by (fast_tac AC_cs 1);
+by (Asm_simp_tac 1);
+by (Fast_tac 1);
val cases = result();
(* ********************************************************************** *)
(* Lemmas used in both cases *)
(* ********************************************************************** *)
goal thy "!!a C. Ord(a) ==> (UN b<a++a. C(b)) = (UN b<a. C(b) Un C(a++b))";
-by (fast_tac (AC_cs addSIs [equalityI] addIs [ltI]
+by (fast_tac (!claset addSIs [equalityI] addIs [ltI]
addSDs [lt_oadd_disj]
addSEs [lt_oadd1, oadd_lt_mono2]) 1);
val UN_oadd = result();
@@ -85,7 +85,7 @@
goalw thy [vv1_def] "vv1(f,m,b) <= f`b";
by (rtac (LetI RS LetI) 1);
by (split_tac [expand_if] 1);
-by (simp_tac (ZF_ss addsimps [domain_uu_subset]) 1);
+by (simp_tac (!simpset addsimps [domain_uu_subset]) 1);
val vv1_subset = result();
(* ********************************************************************** *)
@@ -95,14 +95,14 @@
"!! a f y. [| Ord(a); m:nat |] ==> \
\ (UN b<a++a. gg1(f,a,m)`b) = (UN b<a. f`b)";
by (asm_simp_tac
- (OrdQuant_ss addsimps [UN_oadd, lt_oadd1,
+ (!simpset addsimps [UN_oadd, lt_oadd1,
oadd_le_self RS le_imp_not_lt, lt_Ord,
odiff_oadd_inverse, ltD,
vv1_subset RS Diff_partition, ww1_def]) 1);
val UN_gg1_eq = result();
goal thy "domain(gg1(f,a,m)) = a++a";
-by (simp_tac (ZF_ss addsimps [lam_funtype RS domain_of_fun, gg1_def]) 1);
+by (simp_tac (!simpset addsimps [lam_funtype RS domain_of_fun, gg1_def]) 1);
val domain_gg1 = result();
(* ********************************************************************** *)
@@ -113,7 +113,7 @@
\ ==> P(Least_a, LEAST b. P(Least_a, b))";
by (etac ssubst 1);
by (res_inst_tac [("Q","%z. P(z, LEAST b. P(z, b))")] LeastI2 1);
-by (REPEAT (fast_tac (ZF_cs addSEs [LeastI]) 1));
+by (REPEAT (fast_tac (!claset addSEs [LeastI]) 1));
val nested_LeastI = result();
val nested_Least_instance =
@@ -129,24 +129,24 @@
\ domain(uu(f,b,g,d)) lepoll m); \
\ ALL b<a. f`b lepoll succ(m); b<a++a \
\ |] ==> gg1(f,a,m)`b lepoll m";
-by (asm_simp_tac OrdQuant_ss 1);
-by (safe_tac (OrdQuant_cs addSEs [lt_oadd_odiff_cases]));
+by (Asm_simp_tac 1);
+by (safe_tac (!claset addSEs [lt_oadd_odiff_cases]));
(*Case b<a : show vv1(f,m,b) lepoll m *)
-by (asm_simp_tac (ZF_ss addsimps [vv1_def, Let_def]
+by (asm_simp_tac (!simpset addsimps [vv1_def, Let_def]
setloop split_tac [expand_if]) 1);
-by (fast_tac (AC_cs addIs [nested_Least_instance RS conjunct2]
+by (fast_tac (!claset addIs [nested_Least_instance RS conjunct2]
addSEs [lt_Ord]
addSIs [empty_lepollI]) 1);
(*Case a le b: show ww1(f,m,b--a) lepoll m *)
-by (asm_simp_tac (ZF_ss addsimps [ww1_def]) 1);
+by (asm_simp_tac (!simpset addsimps [ww1_def]) 1);
by (excluded_middle_tac "f`(b--a) = 0" 1);
-by (asm_simp_tac (OrdQuant_ss addsimps [empty_lepollI]) 2);
+by (asm_simp_tac (!simpset addsimps [empty_lepollI]) 2);
by (rtac Diff_lepoll 1);
-by (fast_tac AC_cs 1);
+by (Fast_tac 1);
by (rtac vv1_subset 1);
by (dtac (ospec RS mp) 1);
by (REPEAT (eresolve_tac [asm_rl, oexE] 1));
-by (asm_simp_tac (ZF_ss
+by (asm_simp_tac (!simpset
addsimps [vv1_def, Let_def, lt_Ord,
nested_Least_instance RS conjunct1]) 1);
val gg1_lepoll_m = result();
@@ -162,7 +162,7 @@
goalw thy [uu_def] "!!f. [| b<a; g<a; f`b~=0; f`g~=0; \
\ y*y <= y; (UN b<a. f`b)=y \
\ |] ==> EX d<a. uu(f,b,g,d) ~= 0";
-by (fast_tac (AC_cs addSIs [not_emptyI]
+by (fast_tac (!claset addSIs [not_emptyI]
addSDs [SigmaI RSN (2, subsetD)]
addSEs [not_emptyE]) 1);
val ex_d_uu_not_empty = result();
@@ -171,10 +171,10 @@
\ y*y<=y; (UN b<a. f`b)=y |] \
\ ==> uu(f,b,g,LEAST d. (uu(f,b,g,d) ~= 0)) ~= 0";
by (dtac ex_d_uu_not_empty 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSEs [LeastI, lt_Ord]) 1);
+by (fast_tac (!claset addSEs [LeastI, lt_Ord]) 1);
val uu_not_empty = result();
-goal ZF.thy "!!r. [| r<=A*B; r~=0 |] ==> domain(r)~=0";
+goal upair.thy "!!r. [| r<=A*B; r~=0 |] ==> domain(r)~=0";
by (REPEAT (eresolve_tac [asm_rl, not_emptyE, subsetD RS SigmaE,
sym RSN (2, subst_elem) RS domainI RS not_emptyI] 1));
val not_empty_rel_imp_domain = result();
@@ -188,14 +188,14 @@
THEN (REPEAT (ares_tac [lt_Ord] 1)));
val Least_uu_not_empty_lt_a = result();
-goal ZF.thy "!!B. [| B<=A; a~:B |] ==> B <= A-{a}";
-by (fast_tac ZF_cs 1);
+goal upair.thy "!!B. [| B<=A; a~:B |] ==> B <= A-{a}";
+by (Fast_tac 1);
val subset_Diff_sing = result();
(*Could this be proved more directly?*)
goal thy "!!A B. [| A lepoll m; m lepoll B; B <= A; m:nat |] ==> A=B";
by (etac natE 1);
-by (fast_tac (AC_cs addSDs [lepoll_0_is_0] addSIs [equalityI]) 1);
+by (fast_tac (!claset addSDs [lepoll_0_is_0] addSIs [equalityI]) 1);
by (hyp_subst_tac 1);
by (rtac equalityI 1);
by (assume_tac 2);
@@ -222,7 +222,7 @@
uu_subset1 RSN (4, rel_is_fun)))] 1
THEN TRYALL assume_tac);
by (rtac (eqpoll_sym RS eqpoll_imp_lepoll RSN (2, supset_lepoll_imp_eq)) 1);
-by (REPEAT (fast_tac (AC_cs addSIs [domain_uu_subset, nat_succI]) 1));
+by (REPEAT (fast_tac (!claset addSIs [domain_uu_subset, nat_succI]) 1));
val uu_Least_is_fun = result();
goalw thy [vv2_def]
@@ -232,9 +232,9 @@
\ (UN b<a. f`b)=y; b<a; g<a; m:nat; s:f`b \
\ |] ==> vv2(f,b,g,s) <= f`g";
by (split_tac [expand_if] 1);
-by (fast_tac (FOL_cs addSEs [uu_Least_is_fun]
- addSIs [empty_subsetI, not_emptyI,
- singleton_subsetI, apply_type]) 1);
+by (Step_tac 1);
+be (uu_Least_is_fun RS apply_type) 1;
+by (REPEAT_SOME (fast_tac (!claset addSIs [not_emptyI, singleton_subsetI])));
val vv2_subset = result();
(* ********************************************************************** *)
@@ -248,14 +248,14 @@
\ |] ==> (UN g<a++a. gg2(f,a,b,s) ` g) = y";
by (dtac sym 1);
by (asm_simp_tac
- (OrdQuant_ss addsimps [UN_oadd, lt_oadd1,
+ (!simpset addsimps [UN_oadd, lt_oadd1,
oadd_le_self RS le_imp_not_lt, lt_Ord,
odiff_oadd_inverse, ww2_def,
vv2_subset RS Diff_partition]) 1);
val UN_gg2_eq = result();
goal thy "domain(gg2(f,a,b,s)) = a++a";
-by (simp_tac (ZF_ss addsimps [lam_funtype RS domain_of_fun, gg2_def]) 1);
+by (simp_tac (!simpset addsimps [lam_funtype RS domain_of_fun, gg2_def]) 1);
val domain_gg2 = result();
(* ********************************************************************** *)
@@ -264,9 +264,9 @@
goalw thy [vv2_def]
"!!m. [| m:nat; m~=0 |] ==> vv2(f,b,g,s) lepoll m";
-by (asm_simp_tac (OrdQuant_ss addsimps [empty_lepollI]
+by (asm_simp_tac (!simpset addsimps [empty_lepollI]
setloop split_tac [expand_if]) 1);
-by (fast_tac (AC_cs
+by (fast_tac (!claset
addSDs [le_imp_subset RS subset_imp_lepoll RS lepoll_0_is_0]
addSIs [singleton_eqpoll_1 RS eqpoll_imp_lepoll RS lepoll_trans,
not_lt_imp_le RS le_imp_subset RS subset_imp_lepoll,
@@ -277,11 +277,11 @@
"!!m. [| ALL b<a. f`b lepoll succ(m); g<a; m:nat; vv2(f,b,g,d) <= f`g \
\ |] ==> ww2(f,b,g,d) lepoll m";
by (excluded_middle_tac "f`g = 0" 1);
-by (asm_simp_tac (OrdQuant_ss addsimps [empty_lepollI]) 2);
+by (asm_simp_tac (!simpset addsimps [empty_lepollI]) 2);
by (dtac ospec 1 THEN (assume_tac 1));
by (rtac Diff_lepoll 1
THEN (TRYALL assume_tac));
-by (asm_simp_tac (OrdQuant_ss addsimps [vv2_def, expand_if, not_emptyI]) 1);
+by (asm_simp_tac (!simpset addsimps [vv2_def, expand_if, not_emptyI]) 1);
val ww2_lepoll = result();
goalw thy [gg2_def]
@@ -290,10 +290,10 @@
\ ALL b<a. f`b lepoll succ(m); y*y <= y; \
\ (UN b<a. f`b)=y; b<a; s:f`b; m:nat; m~= 0; g<a++a \
\ |] ==> gg2(f,a,b,s) ` g lepoll m";
-by (asm_simp_tac OrdQuant_ss 1);
-by (safe_tac (OrdQuant_cs addSEs [lt_oadd_odiff_cases, lt_Ord2]));
-by (asm_simp_tac (OrdQuant_ss addsimps [vv2_lepoll]) 1);
-by (asm_simp_tac (ZF_ss addsimps [ww2_lepoll, vv2_subset]) 1);
+by (Asm_simp_tac 1);
+by (safe_tac (!claset addSEs [lt_oadd_odiff_cases, lt_Ord2]));
+by (asm_simp_tac (!simpset addsimps [vv2_lepoll]) 1);
+by (asm_simp_tac (!simpset addsimps [ww2_lepoll, vv2_subset]) 1);
val gg2_lepoll_m = result();
(* ********************************************************************** *)
@@ -305,9 +305,9 @@
by (resolve_tac [quant_domain_uu_lepoll_m RS cases RS disjE] 1
THEN (assume_tac 1));
(* case 1 *)
-by (asm_full_simp_tac (ZF_ss addsimps [lesspoll_succ_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [lesspoll_succ_iff]) 1);
by (res_inst_tac [("x","a++a")] exI 1);
-by (fast_tac (OrdQuant_cs addSIs [Ord_oadd, domain_gg1, UN_gg1_eq,
+by (fast_tac (!claset addSIs [Ord_oadd, domain_gg1, UN_gg1_eq,
gg1_lepoll_m]) 1);
(* case 2 *)
by (REPEAT (eresolve_tac [oexE, conjE] 1));
@@ -318,7 +318,7 @@
by (res_inst_tac [("x","gg2(f,a,b,x)")] exI 1);
(*Calling fast_tac might get rid of the res_inst_tac calls, but it
is just too slow.*)
-by (asm_simp_tac (OrdQuant_ss addsimps
+by (asm_simp_tac (!simpset addsimps
[Ord_oadd, domain_gg2, UN_gg2_eq, gg2_lepoll_m]) 1);
val lemma_ii = result();
@@ -333,14 +333,14 @@
goal thy "ALL n:nat. rec(n, x, %k r. r Un r*r) <= \
\ rec(succ(n), x, %k r. r Un r*r)";
-by (fast_tac (ZF_cs addIs [rec_succ RS ssubst]) 1);
+by (fast_tac (!claset addIs [rec_succ RS ssubst]) 1);
val z_n_subset_z_succ_n = result();
goal thy "!!n. [| ALL n:nat. f(n)<=f(succ(n)); n le m; n : nat; m: nat |] \
\ ==> f(n)<=f(m)";
-by (res_inst_tac [("P","n le m")] impE 1 THEN (REPEAT (assume_tac 2)));
+by (eres_inst_tac [("P","n le m")] rev_mp 1);
by (res_inst_tac [("P","%z. n le z --> f(n) <= f(z)")] nat_induct 1);
-by (REPEAT (fast_tac lt_cs 1));
+by (REPEAT (fast_tac le_cs 1));
val le_subsets = result();
goal thy "!!n m. [| n le m; m:nat |] ==> \
@@ -353,13 +353,14 @@
goal thy "EX y. x Un y*y <= y";
by (res_inst_tac [("x","UN n:nat. rec(n, x, %k r. r Un r*r)")] exI 1);
-by (safe_tac ZF_cs);
-by (fast_tac (ZF_cs addSIs [nat_0I] addss nat_ss) 1);
+by (safe_tac (!claset));
+br (nat_0I RS UN_I) 1;
+by (Asm_simp_tac 1);
by (res_inst_tac [("a","succ(n Un na)")] UN_I 1);
by (eresolve_tac [Un_nat_type RS nat_succI] 1 THEN (assume_tac 1));
by (fast_tac (ZF_cs addIs [le_imp_rec_subset RS subsetD]
addSIs [Un_upper1_le, Un_upper2_le, Un_nat_type]
- addSEs [nat_into_Ord] addss nat_ss) 1);
+ addSEs [nat_into_Ord] addss (!simpset)) 1);
val lemma_iv = result();
(* ********************************************************************** *)
@@ -387,13 +388,13 @@
goal thy "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |] \
\ ==> EX c<a. f`c = {x}";
-by (fast_tac (AC_cs addSEs [lepoll_1_is_sing]) 1);
+by (fast_tac (!claset addSEs [lepoll_1_is_sing]) 1);
val lemma1 = result();
goal thy "!!f. [| (UN b<a. f`b)=y; x:y; ALL b<a. f`b lepoll 1; Ord(a) |] \
\ ==> f` (LEAST i. f`i = {x}) = {x}";
by (dtac lemma1 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSEs [lt_Ord] addIs [LeastI]) 1);
+by (fast_tac (!claset addSEs [lt_Ord] addIs [LeastI]) 1);
val lemma2 = result();
goalw thy [NN_def] "!!y. 1 : NN(y) ==> EX a f. Ord(a) & f:inj(y, a)";
@@ -404,14 +405,14 @@
by (rtac conjI 1 THEN (assume_tac 1));
by (res_inst_tac [("d","%i. THE x. x:f`i")] lam_injective 1);
by (dtac lemma1 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (AC_cs addSEs [Least_le RS lt_trans1 RS ltD, lt_Ord]) 1);
+by (fast_tac (!claset addSEs [Least_le RS lt_trans1 RS ltD, lt_Ord]) 1);
by (resolve_tac [lemma2 RS ssubst] 1 THEN REPEAT (assume_tac 1));
-by (fast_tac (ZF_cs addSIs [the_equality]) 1);
+by (fast_tac (!claset addSIs [the_equality]) 1);
val NN_imp_ex_inj = result();
goal thy "!!y. [| y*y <= y; 1 : NN(y) |] ==> EX r. well_ord(y, r)";
by (dtac NN_imp_ex_inj 1);
-by (fast_tac (ZF_cs addSEs [well_ord_Memrel RSN (2, well_ord_rvimage)]) 1);
+by (fast_tac (!claset addSEs [well_ord_Memrel RSN (2, well_ord_rvimage)]) 1);
val y_well_ord = result();
(* ********************************************************************** *)
@@ -423,10 +424,10 @@
\ ==> n~=0 --> P(n) --> P(1)";
by (res_inst_tac [("n","n")] nat_induct 1);
by (rtac prem1 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (excluded_middle_tac "x=0" 1);
-by (fast_tac ZF_cs 2);
-by (fast_tac (ZF_cs addSIs [prem2]) 1);
+by (Fast_tac 2);
+by (fast_tac (!claset addSIs [prem2]) 1);
val rev_induct_lemma = result();
val prems = goal thy
@@ -453,21 +454,21 @@
(* another helpful lemma *)
goalw thy [NN_def] "!!y. 0:NN(y) ==> y=0";
-by (fast_tac (AC_cs addSIs [equalityI]
+by (fast_tac (!claset addSIs [equalityI]
addSDs [lepoll_0_is_0] addEs [subst]) 1);
val NN_y_0 = result();
goalw thy [WO1_def] "!!Z. WO6 ==> WO1";
by (rtac allI 1);
by (excluded_middle_tac "A=0" 1);
-by (fast_tac (ZF_cs addSIs [well_ord_Memrel, nat_0I RS nat_into_Ord]) 2);
+by (fast_tac (!claset addSIs [well_ord_Memrel, nat_0I RS nat_into_Ord]) 2);
by (res_inst_tac [("x1","A")] (lemma_iv RS revcut_rl) 1);
by (etac exE 1);
by (dtac WO6_imp_NN_not_empty 1);
by (eresolve_tac [Un_subset_iff RS iffD1 RS conjE] 1);
by (eres_inst_tac [("A","NN(y)")] not_emptyE 1);
by (forward_tac [y_well_ord] 1);
-by (fast_tac (ZF_cs addEs [well_ord_subset]) 2);
-by (fast_tac (ZF_cs addSIs [lemma3] addSDs [NN_y_0] addSEs [not_emptyE]) 1);
+by (fast_tac (!claset addEs [well_ord_subset]) 2);
+by (fast_tac (!claset addSIs [lemma3] addSDs [NN_y_0] addSEs [not_emptyE]) 1);
qed "WO6_imp_WO1";
--- a/src/ZF/AC/WO_AC.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO_AC.ML Fri Jan 03 15:01:55 1997 +0100
@@ -9,14 +9,14 @@
goal thy "!!A. [| well_ord(Union(A),r); 0~:A; B:A |] \
\ ==> (THE b. first(b,B,r)) : B";
-by (fast_tac (AC_cs addSEs [well_ord_imp_ex1_first RS theI RS
+by (fast_tac (!claset addSEs [well_ord_imp_ex1_first RS theI RS
(first_def RS def_imp_iff RS iffD1 RS conjunct1)]) 1);
val first_in_B = result();
goal thy "!!A. [| well_ord(Union(A), R); 0~:A |] ==> EX f. f:(PROD X:A. X)";
-by (fast_tac (AC_cs addSEs [first_in_B] addSIs [lam_type]) 1);
+by (fast_tac (!claset addSEs [first_in_B] addSIs [lam_type]) 1);
val ex_choice_fun = result();
goal thy "!!A. well_ord(A, R) ==> EX f. f:(PROD X: Pow(A)-{0}. X)";
-by (fast_tac (AC_cs addSEs [well_ord_subset RS ex_choice_fun]) 1);
+by (fast_tac (!claset addSEs [well_ord_subset RS ex_choice_fun]) 1);
val ex_choice_fun_Pow = result();
--- a/src/ZF/AC/WO_AC.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/WO_AC.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,3 @@
(*Dummy theory to document dependencies *)
-WO_AC = Order + first
+WO_AC = Order
--- a/src/ZF/AC/first.ML Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,29 +0,0 @@
-(* Title: ZF/AC/first.ML
- ID: $Id$
- Author: Krzysztof Grabczewski
-
-Lemmas involving the first element of a well ordered set
-*)
-
-open first;
-
-goalw thy [first_def] "!!b. first(b,B,r) ==> b:B";
-by (fast_tac AC_cs 1);
-val first_is_elem = result();
-
-goalw thy [well_ord_def, wf_on_def, wf_def, first_def]
- "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (EX! b. first(b,B,r))";
-by (REPEAT (eresolve_tac [conjE,allE,disjE] 1));
-by (contr_tac 1);
-by (etac bexE 1);
-by (res_inst_tac [("a","x")] ex1I 1);
-by (fast_tac ZF_cs 2);
-by (rewrite_goals_tac [tot_ord_def, linear_def]);
-by (fast_tac ZF_cs 1);
-val well_ord_imp_ex1_first = result();
-
-goal thy "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B";
-by (dtac well_ord_imp_ex1_first 1 THEN REPEAT (assume_tac 1));
-by (rtac first_is_elem 1);
-by (etac theI 1);
-val the_first_in = result();
--- a/src/ZF/AC/first.thy Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-(* Title: ZF/AC/first.thy
- ID: $Id$
- Author: Krzysztof Grabczewski
-
-Theory helpful in proofs using first element of a well ordered set
-*)
-
-first = Order +
-
-consts
-
- first :: [i, i, i] => o
-
-defs
-
- first_def "first(u, X, R)
- == u:X & (ALL v:X. v~=u --> <u,v> : R)"
-end
--- a/src/ZF/AC/recfunAC16.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/recfunAC16.ML Fri Jan 03 15:01:55 1997 +0100
@@ -33,15 +33,15 @@
(* ********************************************************************** *)
val [prem1, prem2] = goal thy
- "[| !!g r. r <= B(g,r); Ord(i) |] \
+ "[| !!g r. r <= B(g,r); Ord(i) |] \
\ ==> j<i --> transrec2(j, 0, B) <= transrec2(i, 0, B)";
by (resolve_tac [prem2 RS trans_induct] 1);
by (rtac Ord_cases 1 THEN (REPEAT (assume_tac 1)));
-by (fast_tac lt_cs 1);
-by (asm_simp_tac (AC_ss addsimps [transrec2_succ]) 1);
+by (Fast_tac 1);
+by (Asm_simp_tac 1);
by (fast_tac (FOL_cs addSIs [succI1, prem1]
addSEs [ballE, leE, prem1 RSN (2, subset_trans)]) 1);
-by (fast_tac (AC_cs addIs [OUN_I, ltI]
+by (fast_tac (!claset addIs [OUN_I, ltI]
addSEs [Limit_has_succ RS ltE, succI1 RSN (2, Ord_in_Ord) RS le_refl,
transrec2_Limit RS ssubst]) 1);
val transrec2_mono_lemma = result();
@@ -50,7 +50,7 @@
\ ==> transrec2(j, 0, B) <= transrec2(i, 0, B)";
by (resolve_tac [prem2 RS leE] 1);
by (resolve_tac [transrec2_mono_lemma RS impE] 1);
-by (TRYALL (fast_tac (AC_cs addSIs [prem1, prem2, lt_Ord2])));
+by (TRYALL (fast_tac (!claset addSIs [prem1, prem2, lt_Ord2])));
val transrec2_mono = result();
(* ********************************************************************** *)
@@ -60,6 +60,6 @@
goalw thy [recfunAC16_def]
"!!i. i le j ==> recfunAC16(f, g, i, a) <= recfunAC16(f, g, j, a)";
by (rtac transrec2_mono 1);
-by (REPEAT (fast_tac (AC_cs addIs [expand_if RS iffD2]) 1));
+by (REPEAT (fast_tac (!claset addIs [expand_if RS iffD2]) 1));
val recfunAC16_mono = result();
--- a/src/ZF/AC/recfunAC16.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/recfunAC16.thy Fri Jan 03 15:01:55 1997 +0100
@@ -5,15 +5,11 @@
A recursive definition used in the proof of WO2 ==> AC16
*)
-recfunAC16 = Transrec2 + Cardinal +
-
-consts
+recfunAC16 = Cardinal + Epsilon +
- recfunAC16 :: [i, i, i, i] => i
+constdefs
+ recfunAC16 :: [i, i, i, i] => i
-defs
-
- recfunAC16_def
"recfunAC16(f,fa,i,a) ==
transrec2(i, 0,
%g r. if(EX y:r. fa`g <= y, r,
--- a/src/ZF/AC/rel_is_fun.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/rel_is_fun.ML Fri Jan 03 15:01:55 1997 +0100
@@ -15,36 +15,28 @@
by (res_inst_tac [("x",
"lam x:domain(u). LEAST i. EX y. <x,y> : u & f`<x,y> = i")] exI 1);
by (res_inst_tac [("d","%y. fst(converse(f)`y)")] lam_injective 1);
-by (fast_tac (ZF_cs addIs [LeastI2, nat_into_Ord RS Ord_in_Ord,
+by (fast_tac (!claset addIs [LeastI2, nat_into_Ord RS Ord_in_Ord,
inj_is_fun RS apply_type]) 1);
by (etac domainE 1);
by (forward_tac [inj_is_fun RS apply_type] 1 THEN (atac 1));
by (rtac LeastI2 1);
-by (REPEAT (fast_tac (ZF_cs addIs [fst_conv, left_inverse RS ssubst]
+by (REPEAT (fast_tac (!claset addIs [fst_conv, left_inverse RS ssubst]
addSEs [nat_into_Ord RS Ord_in_Ord]) 1));
val lepoll_m_imp_domain_lepoll_m = result();
-goal ZF.thy "!!r. [| <a,c> : r; c~=b |] ==> domain(r-{<a,b>}) = domain(r)";
-by (rtac equalityI 1);
-by (fast_tac (ZF_cs addSIs [domain_mono]) 1);
-by (rtac subsetI 1);
-by (excluded_middle_tac "x = a" 1);
-by (REPEAT (fast_tac ZF_cs 1));
-val domain_Diff_eq_domain = result();
-
goalw Cardinal.thy [function_def]
"!!r. [| succ(m) lepoll domain(r); r lepoll succ(m); m:nat |] ==> \
\ function(r)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (resolve_tac [excluded_middle RS disjE] 1 THEN (atac 2));
-by (fast_tac (ZF_cs addSEs [lepoll_trans RS succ_lepoll_natE,
+by (fast_tac (!claset addSEs [lepoll_trans RS succ_lepoll_natE,
Diff_sing_lepoll RSN (2, lepoll_m_imp_domain_lepoll_m)]
- addEs [not_sym RSN (2, domain_Diff_eq_domain) RS subst]) 1);
+ addEs [not_sym RSN (2, domain_Diff_eq) RS subst]) 1);
val rel_domain_ex1 = result();
goal Cardinal.thy
"!!r. [| succ(m) lepoll domain(r); r lepoll succ(m); m:nat; \
\ r<=A*B; A=domain(r) |] ==> r: A->B";
by (hyp_subst_tac 1);
-by (asm_simp_tac (ZF_ss addsimps [Pi_iff, rel_domain_ex1]) 1);
+by (asm_simp_tac (!simpset addsimps [Pi_iff, rel_domain_ex1]) 1);
val rel_is_fun = result();
--- a/src/ZF/AC/rel_is_fun.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/AC/rel_is_fun.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,1 @@
-(*Dummy theory to document dependencies *)
-
-rel_is_fun = "Cardinal"
+rel_is_fun = Cardinal
--- a/src/ZF/Arith.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Arith.ML Fri Jan 03 15:01:55 1997 +0100
@@ -29,9 +29,11 @@
goal Arith.thy "rec(succ(m),a,b) = b(m, rec(m,a,b))";
by (rtac rec_trans 1);
-by (simp_tac (ZF_ss addsimps [nat_case_succ, nat_succI]) 1);
+by (Simp_tac 1);
qed "rec_succ";
+Addsimps [rec_0, rec_succ];
+
val major::prems = goal Arith.thy
"[| n: nat; \
\ a: C(0); \
@@ -39,22 +41,17 @@
\ |] ==> rec(n,a,b) : C(n)";
by (rtac (major RS nat_induct) 1);
by (ALLGOALS
- (asm_simp_tac (ZF_ss addsimps (prems@[rec_0,rec_succ]))));
+ (asm_simp_tac (!simpset addsimps prems)));
qed "rec_type";
+Addsimps [rec_type, nat_0_le, nat_le_refl];
val nat_typechecks = [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat];
-val nat_simps = [rec_0, rec_succ, not_lt0, nat_0_le, le0_iff, succ_le_iff,
- nat_le_refl];
-
-val nat_ss = ZF_ss addsimps (nat_simps @ nat_typechecks);
-
-
goal Arith.thy "!!k. [| 0<k; k: nat |] ==> EX j: nat. k = succ(j)";
by (etac rev_mp 1);
by (etac nat_induct 1);
-by (simp_tac nat_ss 1);
-by (fast_tac ZF_cs 1);
+by (Simp_tac 1);
+by (Fast_tac 1);
val lemma = result();
(* [| 0 < k; k: nat; !!j. [| j: nat; k = succ(j) |] ==> Q |] ==> Q *)
@@ -75,6 +72,8 @@
"succ(m) #+ n = succ(m #+ n)"
(fn _=> [ (rtac rec_succ 1) ]);
+Addsimps [add_0, add_succ];
+
(** Multiplication **)
qed_goalw "mult_type" Arith.thy [mult_def]
@@ -90,6 +89,8 @@
"succ(m) #* n = n #+ (m #* n)"
(fn _ => [ (rtac rec_succ 1) ]);
+Addsimps [mult_0, mult_succ];
+
(** Difference **)
qed_goalw "diff_type" Arith.thy [diff_def]
@@ -104,35 +105,32 @@
"n:nat ==> 0 #- n = 0"
(fn [prem]=>
[ (rtac (prem RS nat_induct) 1),
- (ALLGOALS (asm_simp_tac nat_ss)) ]);
+ (ALLGOALS (Asm_simp_tac)) ]);
(*Must simplify BEFORE the induction!! (Else we get a critical pair)
succ(m) #- succ(n) rewrites to pred(succ(m) #- n) *)
qed_goalw "diff_succ_succ" Arith.thy [diff_def]
"[| m:nat; n:nat |] ==> succ(m) #- succ(n) = m #- n"
(fn prems=>
- [ (asm_simp_tac (nat_ss addsimps prems) 1),
+ [ (asm_simp_tac (!simpset addsimps prems) 1),
(nat_ind_tac "n" prems 1),
- (ALLGOALS (asm_simp_tac (nat_ss addsimps prems))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps prems))) ]);
+
+Addsimps [diff_0, diff_0_eq_0, diff_succ_succ];
val prems = goal Arith.thy
"[| m:nat; n:nat |] ==> m #- n le m";
by (rtac (prems MRS diff_induct) 1);
by (etac leE 3);
by (ALLGOALS
- (asm_simp_tac
- (nat_ss addsimps (prems @ [le_iff, diff_0, diff_0_eq_0,
- diff_succ_succ, nat_into_Ord]))));
+ (asm_simp_tac (!simpset addsimps (prems @ [le_iff, nat_into_Ord]))));
qed "diff_le_self";
(*** Simplification over add, mult, diff ***)
val arith_typechecks = [add_type, mult_type, diff_type];
-val arith_simps = [add_0, add_succ,
- mult_0, mult_succ,
- diff_0, diff_0_eq_0, diff_succ_succ];
+Addsimps arith_typechecks;
-val arith_ss = nat_ss addsimps (arith_simps@arith_typechecks);
(*** Addition ***)
@@ -141,7 +139,7 @@
"m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
- (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps prems))) ]);
(*The following two lemmas are used for add_commute and sometimes
elsewhere, since they are safe for rewriting.*)
@@ -149,26 +147,27 @@
"m:nat ==> m #+ 0 = m"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
- (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps prems))) ]);
qed_goal "add_succ_right" Arith.thy
"m:nat ==> m #+ succ(n) = succ(m #+ n)"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
- (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps prems))) ]);
+
+Addsimps [add_0_right, add_succ_right];
(*Commutative law for addition*)
qed_goal "add_commute" Arith.thy
"!!m n. [| m:nat; n:nat |] ==> m #+ n = n #+ m"
(fn _ =>
[ (nat_ind_tac "n" [] 1),
- (ALLGOALS
- (asm_simp_tac (arith_ss addsimps [add_0_right, add_succ_right]))) ]);
+ (ALLGOALS Asm_simp_tac) ]);
(*for a/c rewriting*)
qed_goal "add_left_commute" Arith.thy
"!!m n k. [| m:nat; n:nat |] ==> m#+(n#+k)=n#+(m#+k)"
- (fn _ => [asm_simp_tac (ZF_ss addsimps [add_assoc RS sym, add_commute]) 1]);
+ (fn _ => [asm_simp_tac(!simpset addsimps [add_assoc RS sym, add_commute]) 1]);
(*Addition is an AC-operator*)
val add_ac = [add_assoc, add_commute, add_left_commute];
@@ -178,7 +177,7 @@
"[| k #+ m = k #+ n; k:nat |] ==> m=n";
by (rtac (eqn RS rev_mp) 1);
by (nat_ind_tac "k" [knat] 1);
-by (ALLGOALS (simp_tac arith_ss));
+by (ALLGOALS (Simp_tac));
qed "add_left_cancel";
(*** Multiplication ***)
@@ -188,53 +187,53 @@
"!!m. m:nat ==> m #* 0 = 0"
(fn _=>
[ (nat_ind_tac "m" [] 1),
- (ALLGOALS (asm_simp_tac arith_ss)) ]);
+ (ALLGOALS (Asm_simp_tac)) ]);
(*right successor law for multiplication*)
qed_goal "mult_succ_right" Arith.thy
"!!m n. [| m:nat; n:nat |] ==> m #* succ(n) = m #+ (m #* n)"
(fn _ =>
[ (nat_ind_tac "m" [] 1),
- (ALLGOALS (asm_simp_tac (arith_ss addsimps add_ac))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps add_ac))) ]);
+
+Addsimps [mult_0_right, mult_succ_right];
goal Arith.thy "!!n. n:nat ==> 1 #* n = n";
-by (asm_simp_tac (arith_ss addsimps [add_0_right]) 1);
+by (Asm_simp_tac 1);
qed "mult_1";
goal Arith.thy "!!n. n:nat ==> n #* 1 = n";
-by (asm_simp_tac (arith_ss addsimps [mult_0_right, add_0_right,
- mult_succ_right]) 1);
+by (Asm_simp_tac 1);
qed "mult_1_right";
(*Commutative law for multiplication*)
qed_goal "mult_commute" Arith.thy
- "[| m:nat; n:nat |] ==> m #* n = n #* m"
- (fn prems=>
- [ (nat_ind_tac "m" prems 1),
- (ALLGOALS (asm_simp_tac
- (arith_ss addsimps (prems@[mult_0_right, mult_succ_right])))) ]);
+ "!!m n. [| m:nat; n:nat |] ==> m #* n = n #* m"
+ (fn _=>
+ [ (nat_ind_tac "m" [] 1),
+ (ALLGOALS Asm_simp_tac) ]);
(*addition distributes over multiplication*)
qed_goal "add_mult_distrib" Arith.thy
"!!m n. [| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)"
(fn _=>
[ (etac nat_induct 1),
- (ALLGOALS (asm_simp_tac (arith_ss addsimps [add_assoc RS sym]))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps [add_assoc RS sym]))) ]);
(*Distributive law on the left; requires an extra typing premise*)
qed_goal "add_mult_distrib_left" Arith.thy
"!!m. [| m:nat; n:nat; k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)"
(fn prems=>
[ (nat_ind_tac "m" [] 1),
- (asm_simp_tac (arith_ss addsimps [mult_0_right]) 1),
- (asm_simp_tac (arith_ss addsimps ([mult_succ_right] @ add_ac)) 1) ]);
+ (Asm_simp_tac 1),
+ (asm_simp_tac (!simpset addsimps add_ac) 1) ]);
(*Associative law for multiplication*)
qed_goal "mult_assoc" Arith.thy
"!!m n k. [| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)"
(fn _=>
[ (etac nat_induct 1),
- (ALLGOALS (asm_simp_tac (arith_ss addsimps [add_mult_distrib]))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps [add_mult_distrib]))) ]);
(*for a/c rewriting*)
qed_goal "mult_left_commute" Arith.thy
@@ -253,7 +252,7 @@
"m:nat ==> m #- m = 0"
(fn prems=>
[ (nat_ind_tac "m" prems 1),
- (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]);
+ (ALLGOALS (asm_simp_tac (!simpset addsimps prems))) ]);
(*Addition is the inverse of subtraction*)
goal Arith.thy "!!m n. [| n le m; m:nat |] ==> n #+ (m#-n) = m";
@@ -261,7 +260,7 @@
by (etac nat_succI 1);
by (etac rev_mp 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
-by (ALLGOALS (asm_simp_tac arith_ss));
+by (ALLGOALS Asm_simp_tac);
qed "add_diff_inverse";
(*Proof is IDENTICAL to that above*)
@@ -270,7 +269,7 @@
by (etac nat_succI 1);
by (etac rev_mp 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
-by (ALLGOALS (asm_simp_tac arith_ss));
+by (ALLGOALS Asm_simp_tac);
qed "diff_succ";
(** Subtraction is the inverse of addition. **)
@@ -278,7 +277,7 @@
val [mnat,nnat] = goal Arith.thy
"[| m:nat; n:nat |] ==> (n#+m) #- n = m";
by (rtac (nnat RS nat_induct) 1);
-by (ALLGOALS (asm_simp_tac (arith_ss addsimps [mnat])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [mnat])));
qed "diff_add_inverse";
goal Arith.thy
@@ -290,19 +289,19 @@
goal Arith.thy
"!!k. [| k:nat; m: nat; n: nat |] ==> (k#+m) #- (k#+n) = m #- n";
by (nat_ind_tac "k" [] 1);
-by (ALLGOALS (asm_simp_tac arith_ss));
+by (ALLGOALS Asm_simp_tac);
qed "diff_cancel";
goal Arith.thy
"!!n. [| k:nat; m: nat; n: nat |] ==> (m#+k) #- (n#+k) = m #- n";
val add_commute_k = read_instantiate [("n","k")] add_commute;
-by (asm_simp_tac (arith_ss addsimps [add_commute_k, diff_cancel]) 1);
+by (asm_simp_tac (!simpset addsimps [add_commute_k, diff_cancel]) 1);
qed "diff_cancel2";
val [mnat,nnat] = goal Arith.thy
"[| m:nat; n:nat |] ==> n #- (n#+m) = 0";
by (rtac (nnat RS nat_induct) 1);
-by (ALLGOALS (asm_simp_tac (arith_ss addsimps [mnat])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [mnat])));
qed "diff_add_0";
(** Difference distributes over multiplication **)
@@ -310,13 +309,13 @@
goal Arith.thy
"!!m n. [| m:nat; n: nat; k:nat |] ==> (m #- n) #* k = (m #* k) #- (n #* k)";
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
-by (ALLGOALS (asm_simp_tac (arith_ss addsimps [diff_cancel])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [diff_cancel])));
qed "diff_mult_distrib" ;
goal Arith.thy
"!!m. [| m:nat; n: nat; k:nat |] ==> k #* (m #- n) = (k #* m) #- (k #* n)";
val mult_commute_k = read_instantiate [("m","k")] mult_commute;
-by (asm_simp_tac (arith_ss addsimps
+by (asm_simp_tac (!simpset addsimps
[mult_commute_k, diff_mult_distrib]) 1);
qed "diff_mult_distrib2" ;
@@ -327,7 +326,7 @@
by (etac rev_mp 1);
by (etac rev_mp 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
-by (ALLGOALS (asm_simp_tac (nat_ss addsimps [diff_le_self,diff_succ_succ])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [diff_le_self,diff_succ_succ])));
qed "div_termination";
val div_rls = (*for mod and div*)
@@ -335,8 +334,8 @@
[Ord_transrec_type, apply_type, div_termination RS ltD, if_type,
nat_into_Ord, not_lt_iff_le RS iffD1];
-val div_ss = ZF_ss addsimps [nat_into_Ord, div_termination RS ltD,
- not_lt_iff_le RS iffD2];
+val div_ss = (!simpset) addsimps [nat_into_Ord, div_termination RS ltD,
+ not_lt_iff_le RS iffD2];
(*Type checking depends upon termination!*)
goalw Arith.thy [mod_def] "!!m n. [| 0<n; m:nat; n:nat |] ==> m mod n : nat";
@@ -354,6 +353,8 @@
by (asm_simp_tac div_ss 1);
qed "mod_geq";
+Addsimps [mod_type, mod_less, mod_geq];
+
(*** Quotient ***)
(*Type checking depends upon termination!*)
@@ -374,17 +375,18 @@
by (asm_simp_tac div_ss 1);
qed "div_geq";
+Addsimps [div_type, div_less, div_geq];
+
(*A key result*)
goal Arith.thy
"!!m n. [| 0<n; m:nat; n:nat |] ==> (m div n)#*n #+ m mod n = m";
by (etac complete_induct 1);
by (excluded_middle_tac "x<n" 1);
(*case x<n*)
-by (asm_simp_tac (arith_ss addsimps [mod_less, div_less]) 2);
+by (Asm_simp_tac 2);
(*case n le x*)
by (asm_full_simp_tac
- (arith_ss addsimps [not_lt_iff_le, nat_into_Ord,
- mod_geq, div_geq, add_assoc,
+ (!simpset addsimps [not_lt_iff_le, nat_into_Ord, add_assoc,
div_termination RS ltD, add_diff_inverse]) 1);
qed "mod_div_equality";
@@ -396,26 +398,26 @@
by (etac complete_induct 1);
by (excluded_middle_tac "succ(x)<n" 1);
(* case succ(x) < n *)
-by (asm_simp_tac (arith_ss addsimps [mod_less, nat_le_refl RS lt_trans,
+by (asm_simp_tac (!simpset addsimps [mod_less, nat_le_refl RS lt_trans,
succ_neq_self]) 2);
-by (asm_simp_tac (arith_ss addsimps [ltD RS mem_imp_not_eq]) 2);
+by (asm_simp_tac (!simpset addsimps [ltD RS mem_imp_not_eq]) 2);
(* case n le succ(x) *)
by (asm_full_simp_tac
- (arith_ss addsimps [not_lt_iff_le, nat_into_Ord, mod_geq]) 1);
+ (!simpset addsimps [not_lt_iff_le, nat_into_Ord, mod_geq]) 1);
by (etac leE 1);
-by (asm_simp_tac (arith_ss addsimps [div_termination RS ltD, diff_succ,
+by (asm_simp_tac (!simpset addsimps [div_termination RS ltD, diff_succ,
mod_geq]) 1);
-by (asm_simp_tac (arith_ss addsimps [mod_less, diff_self_eq_0]) 1);
+by (asm_simp_tac (!simpset addsimps [mod_less, diff_self_eq_0]) 1);
qed "mod_succ";
goal Arith.thy "!!m n. [| 0<n; m:nat; n:nat |] ==> m mod n < n";
by (etac complete_induct 1);
by (excluded_middle_tac "x<n" 1);
(*case x<n*)
-by (asm_simp_tac (arith_ss addsimps [mod_less]) 2);
+by (asm_simp_tac (!simpset addsimps [mod_less]) 2);
(*case n le x*)
by (asm_full_simp_tac
- (arith_ss addsimps [not_lt_iff_le, nat_into_Ord,
+ (!simpset addsimps [not_lt_iff_le, nat_into_Ord,
mod_geq, div_termination RS ltD]) 1);
qed "mod_less_divisor";
@@ -423,22 +425,22 @@
goal Arith.thy
"!!k b. [| k: nat; b<2 |] ==> k mod 2 = b | k mod 2 = if(b=1,0,1)";
by (subgoal_tac "k mod 2: 2" 1);
-by (asm_simp_tac (arith_ss addsimps [mod_less_divisor RS ltD]) 2);
+by (asm_simp_tac (!simpset addsimps [mod_less_divisor RS ltD]) 2);
by (dtac ltD 1);
-by (asm_simp_tac (ZF_ss setloop split_tac [expand_if]) 1);
-by (fast_tac ZF_cs 1);
+by (asm_simp_tac (!simpset setloop split_tac [expand_if]) 1);
+by (Fast_tac 1);
qed "mod2_cases";
goal Arith.thy "!!m. m:nat ==> succ(succ(m)) mod 2 = m mod 2";
by (subgoal_tac "m mod 2: 2" 1);
-by (asm_simp_tac (arith_ss addsimps [mod_less_divisor RS ltD]) 2);
-by (asm_simp_tac (arith_ss addsimps [mod_succ] setloop step_tac ZF_cs) 1);
+by (asm_simp_tac (!simpset addsimps [mod_less_divisor RS ltD]) 2);
+by (asm_simp_tac (!simpset addsimps [mod_succ] setloop step_tac (!claset)) 1);
qed "mod2_succ_succ";
goal Arith.thy "!!m. m:nat ==> (m#+m) mod 2 = 0";
by (etac nat_induct 1);
-by (simp_tac (arith_ss addsimps [mod_less]) 1);
-by (asm_simp_tac (arith_ss addsimps [mod2_succ_succ, add_succ_right]) 1);
+by (simp_tac (!simpset addsimps [mod_less]) 1);
+by (asm_simp_tac (!simpset addsimps [mod2_succ_succ, add_succ_right]) 1);
qed "mod2_add_self";
@@ -446,7 +448,7 @@
goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le m #+ n";
by (etac nat_induct 1);
-by (ALLGOALS (asm_simp_tac arith_ss));
+by (ALLGOALS Asm_simp_tac);
qed "add_le_self";
goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le n #+ m";
@@ -461,7 +463,7 @@
by (forward_tac [lt_nat_in_nat] 1);
by (assume_tac 1);
by (etac succ_lt_induct 1);
-by (ALLGOALS (asm_simp_tac (arith_ss addsimps [leI])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [leI])));
qed "add_lt_mono1";
(*strict, in both arguments*)
@@ -481,7 +483,7 @@
\ i le j; j:k \
\ |] ==> f(i) le f(j)";
by (cut_facts_tac prems 1);
-by (fast_tac (lt_cs addSIs [lt_mono,ford] addSEs [leE]) 1);
+by (fast_tac (le_cs addSIs [lt_mono,ford] addSEs [leE]) 1);
qed "Ord_lt_mono_imp_le_mono";
(*le monotonicity, 1st argument*)
@@ -507,8 +509,7 @@
goal Arith.thy "!!i j k. [| i le j; j:nat; k:nat |] ==> i#*k le j#*k";
by (forward_tac [lt_nat_in_nat] 1);
by (nat_ind_tac "k" [] 2);
-by (ALLGOALS (asm_simp_tac (arith_ss addsimps [mult_0_right, mult_succ_right,
- add_le_mono])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [add_le_mono])));
qed "mult_le_mono1";
(* le monotonicity, BOTH arguments*)
@@ -526,17 +527,17 @@
goal Arith.thy "!!i j k. [| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j";
by (etac zero_lt_natE 1);
by (forward_tac [lt_nat_in_nat] 2);
-by (ALLGOALS (asm_simp_tac arith_ss));
+by (ALLGOALS Asm_simp_tac);
by (nat_ind_tac "x" [] 1);
-by (ALLGOALS (asm_simp_tac (arith_ss addsimps [add_0_right, add_lt_mono])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [add_lt_mono])));
qed "mult_lt_mono2";
goal Arith.thy "!!k. [| m: nat; n: nat |] ==> 0 < m#*n <-> 0<m & 0<n";
-by (best_tac (ZF_cs addEs [natE] addss (arith_ss addsimps [mult_0_right])) 1);
+by (best_tac (!claset addEs [natE] addss (!simpset)) 1);
qed "zero_lt_mult_iff";
goal Arith.thy "!!k. [| m: nat; n: nat |] ==> m#*n = 1 <-> m=1 & n=1";
-by (best_tac (ZF_cs addEs [natE] addss (arith_ss addsimps [mult_0_right])) 1);
+by (best_tac (!claset addEs [natE] addss (!simpset)) 1);
qed "mult_eq_1_iff";
(*Cancellation law for division*)
@@ -544,10 +545,10 @@
"!!k. [| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n";
by (eres_inst_tac [("i","m")] complete_induct 1);
by (excluded_middle_tac "x<n" 1);
-by (asm_simp_tac (arith_ss addsimps [div_less, zero_lt_mult_iff,
+by (asm_simp_tac (!simpset addsimps [div_less, zero_lt_mult_iff,
mult_lt_mono2]) 2);
by (asm_full_simp_tac
- (arith_ss addsimps [not_lt_iff_le, nat_into_Ord,
+ (!simpset addsimps [not_lt_iff_le, nat_into_Ord,
zero_lt_mult_iff, le_refl RS mult_le_mono, div_geq,
diff_mult_distrib2 RS sym,
div_termination RS ltD]) 1);
@@ -558,10 +559,10 @@
\ (k#*m) mod (k#*n) = k #* (m mod n)";
by (eres_inst_tac [("i","m")] complete_induct 1);
by (excluded_middle_tac "x<n" 1);
-by (asm_simp_tac (arith_ss addsimps [mod_less, zero_lt_mult_iff,
+by (asm_simp_tac (!simpset addsimps [mod_less, zero_lt_mult_iff,
mult_lt_mono2]) 2);
by (asm_full_simp_tac
- (arith_ss addsimps [not_lt_iff_le, nat_into_Ord,
+ (!simpset addsimps [not_lt_iff_le, nat_into_Ord,
zero_lt_mult_iff, le_refl RS mult_le_mono, mod_geq,
diff_mult_distrib2 RS sym,
div_termination RS ltD]) 1);
@@ -575,13 +576,24 @@
by (rtac disjCI 1);
by (dtac sym 1);
by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I]));
-by (fast_tac (lt_cs addss (arith_ss addsimps [mult_0_right])) 1);
-by (fast_tac (lt_cs addDs [mono_lemma]
- addss (arith_ss addsimps [mult_1_right])) 1);
+by (fast_tac (!claset addss (!simpset)) 1);
+by (fast_tac (le_cs addDs [mono_lemma]
+ addss (!simpset addsimps [mult_1_right])) 1);
qed "mult_eq_self_implies_10";
-val arith_ss0 = arith_ss
-and arith_ss = arith_ss addsimps [add_0_right, add_succ_right,
- mult_0_right, mult_succ_right,
- mod_type, div_type,
- mod_less, mod_geq, div_less, div_geq];
+
+(*Thanks to Sten Agerholm*)
+goal Arith.thy (* add_le_elim1 *)
+ "!!m n k. [|m#+n le m#+k; m:nat; n:nat; k:nat|] ==> n le k";
+be rev_mp 1;
+by (eres_inst_tac[("n","n")]nat_induct 1);
+by (Asm_simp_tac 1);
+by (Step_tac 1);
+by (asm_full_simp_tac (!simpset addsimps [not_le_iff_lt,nat_into_Ord]) 1);
+by (etac lt_asym 1);
+by (assume_tac 1);
+by (Asm_full_simp_tac 1);
+by (asm_full_simp_tac (!simpset addsimps [le_iff, nat_into_Ord]) 1);
+by (Fast_tac 1);
+qed "add_le_elim1";
+
--- a/src/ZF/Arith.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Arith.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
Arithmetic operators and their definitions
*)
-Arith = Epsilon + "simpdata" +
+Arith = Epsilon +
consts
rec :: [i, i, [i,i]=>i]=>i
"#*" :: [i,i]=>i (infixl 70)
--- a/src/ZF/Bool.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Bool.ML Fri Jan 03 15:01:55 1997 +0100
@@ -10,6 +10,10 @@
val bool_defs = [bool_def,cond_def];
+goalw Bool.thy [succ_def] "{0} = 1";
+br refl 1;
+qed "singleton_0";
+
(* Introduction rules *)
goalw Bool.thy bool_defs "1 : bool";
@@ -45,11 +49,14 @@
by (rtac (succ_not_0 RS not_sym RS if_not_P) 1);
qed "cond_0";
-val major::prems = goal Bool.thy
- "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)";
-by (rtac (major RS boolE) 1);
-by (asm_simp_tac (ZF_ss addsimps (cond_1::prems)) 1);
-by (asm_simp_tac (ZF_ss addsimps (cond_0::prems)) 1);
+Addsimps [cond_1, cond_0];
+
+fun bool_tac i = fast_tac (!claset addSEs [boolE] addss (!simpset)) i;
+
+
+goal Bool.thy
+ "!!b. [| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)";
+by (bool_tac 1);
qed "cond_type";
val [rew] = goal Bool.thy "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c";
@@ -65,12 +72,11 @@
fun conds def = [standard (def RS def_cond_1), standard (def RS def_cond_0)];
val [not_1,not_0] = conds not_def;
-
val [and_1,and_0] = conds and_def;
+val [or_1,or_0] = conds or_def;
+val [xor_1,xor_0] = conds xor_def;
-val [or_1,or_0] = conds or_def;
-
-val [xor_1,xor_0] = conds xor_def;
+Addsimps [not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0];
qed_goalw "not_type" Bool.thy [not_def]
"a:bool ==> not(a) : bool"
@@ -89,71 +95,65 @@
(fn prems=> [ (typechk_tac(prems@[bool_1I, bool_0I, cond_type, not_type])) ]);
val bool_typechecks = [bool_1I, bool_0I, cond_type, not_type, and_type,
- or_type, xor_type]
-
-val bool_simps = [cond_1,cond_0,not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0];
-
-val bool_ss0 = ZF_ss addsimps bool_simps;
-
-fun bool0_tac i =
- EVERY [etac boolE i, asm_simp_tac bool_ss0 (i+1), asm_simp_tac bool_ss0 i];
+ or_type, xor_type];
(*** Laws for 'not' ***)
goal Bool.thy "!!a. a:bool ==> not(not(a)) = a";
-by (bool0_tac 1);
+by (bool_tac 1);
qed "not_not";
goal Bool.thy "!!a b. a:bool ==> not(a and b) = not(a) or not(b)";
-by (bool0_tac 1);
+by (bool_tac 1);
qed "not_and";
goal Bool.thy "!!a b. a:bool ==> not(a or b) = not(a) and not(b)";
-by (bool0_tac 1);
+by (bool_tac 1);
qed "not_or";
+Addsimps [not_not, not_and, not_or];
+
(*** Laws about 'and' ***)
goal Bool.thy "!!a. a: bool ==> a and a = a";
-by (bool0_tac 1);
+by (bool_tac 1);
qed "and_absorb";
+Addsimps [and_absorb];
+
goal Bool.thy "!!a b. [| a: bool; b:bool |] ==> a and b = b and a";
-by (etac boolE 1);
-by (bool0_tac 1);
-by (bool0_tac 1);
+by (bool_tac 1);
qed "and_commute";
-goal Bool.thy
- "!!a. a: bool ==> (a and b) and c = a and (b and c)";
-by (bool0_tac 1);
+goal Bool.thy "!!a. a: bool ==> (a and b) and c = a and (b and c)";
+by (bool_tac 1);
qed "and_assoc";
goal Bool.thy
"!!a. [| a: bool; b:bool; c:bool |] ==> \
\ (a or b) and c = (a and c) or (b and c)";
-by (REPEAT (bool0_tac 1));
+by (bool_tac 1);
qed "and_or_distrib";
(** binary orion **)
goal Bool.thy "!!a. a: bool ==> a or a = a";
-by (bool0_tac 1);
+by (bool_tac 1);
qed "or_absorb";
+Addsimps [or_absorb];
+
goal Bool.thy "!!a b. [| a: bool; b:bool |] ==> a or b = b or a";
-by (etac boolE 1);
-by (bool0_tac 1);
-by (bool0_tac 1);
+by (bool_tac 1);
qed "or_commute";
goal Bool.thy "!!a. a: bool ==> (a or b) or c = a or (b or c)";
-by (bool0_tac 1);
+by (bool_tac 1);
qed "or_assoc";
goal Bool.thy
"!!a b c. [| a: bool; b: bool; c: bool |] ==> \
\ (a and b) or c = (a or c) and (b or c)";
-by (REPEAT (bool0_tac 1));
+by (bool_tac 1);
qed "or_and_distrib";
--- a/src/ZF/Bool.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Bool.thy Fri Jan 03 15:01:55 1997 +0100
@@ -8,7 +8,7 @@
2 is equal to bool, but serves a different purpose
*)
-Bool = ZF + "simpdata" +
+Bool = upair +
consts
"1" :: i ("1")
"2" :: i ("2")
--- a/src/ZF/Cardinal.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Cardinal.ML Fri Jan 03 15:01:55 1997 +0100
@@ -25,7 +25,7 @@
\ X - lfp(X, %W. X - g``(Y - f``W)) ";
by (res_inst_tac [("P", "%u. ?v = X-u")]
(decomp_bnd_mono RS lfp_Tarski RS ssubst) 1);
-by (simp_tac (ZF_ss addsimps [subset_refl, double_complement,
+by (simp_tac (!simpset addsimps [subset_refl, double_complement,
gfun RS fun_is_rel RS image_subset]) 1);
qed "Banach_last_equation";
@@ -45,7 +45,7 @@
"[| f: inj(X,Y); g: inj(Y,X) |] ==> EX h. h: bij(X,Y)";
by (cut_facts_tac prems 1);
by (cut_facts_tac [(prems RL [inj_is_fun]) MRS decomposition] 1);
-by (fast_tac (ZF_cs addSIs [restrict_bij,bij_disjoint_Un]
+by (fast_tac (!claset addSIs [restrict_bij,bij_disjoint_Un]
addIs [bij_converse_bij]) 1);
(* The instantiation of exI to "restrict(f,XA) Un converse(restrict(g,YB))"
is forced by the context!! *)
@@ -62,12 +62,12 @@
bind_thm ("eqpoll_refl", id_bij RS bij_imp_eqpoll);
goalw Cardinal.thy [eqpoll_def] "!!X Y. X eqpoll Y ==> Y eqpoll X";
-by (fast_tac (ZF_cs addIs [bij_converse_bij]) 1);
+by (fast_tac (!claset addIs [bij_converse_bij]) 1);
qed "eqpoll_sym";
goalw Cardinal.thy [eqpoll_def]
"!!X Y. [| X eqpoll Y; Y eqpoll Z |] ==> X eqpoll Z";
-by (fast_tac (ZF_cs addIs [comp_bij]) 1);
+by (fast_tac (!claset addIs [comp_bij]) 1);
qed "eqpoll_trans";
(** Le-pollence is a partial ordering **)
@@ -83,12 +83,12 @@
goalw Cardinal.thy [eqpoll_def, bij_def, lepoll_def]
"!!X Y. X eqpoll Y ==> X lepoll Y";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "eqpoll_imp_lepoll";
goalw Cardinal.thy [lepoll_def]
"!!X Y. [| X lepoll Y; Y lepoll Z |] ==> X lepoll Z";
-by (fast_tac (ZF_cs addIs [comp_inj]) 1);
+by (fast_tac (!claset addIs [comp_inj]) 1);
qed "lepoll_trans";
(*Asymmetry law*)
@@ -106,7 +106,7 @@
qed "eqpollE";
goal Cardinal.thy "X eqpoll Y <-> X lepoll Y & Y lepoll X";
-by (fast_tac (ZF_cs addIs [eqpollI] addSEs [eqpollE]) 1);
+by (fast_tac (!claset addIs [eqpollI] addSEs [eqpollE]) 1);
qed "eqpoll_iff";
goalw Cardinal.thy [lepoll_def, inj_def] "!!A. A lepoll 0 ==> A = 0";
@@ -117,70 +117,70 @@
bind_thm ("empty_lepollI", empty_subsetI RS subset_imp_lepoll);
goal Cardinal.thy "A lepoll 0 <-> A=0";
-by (fast_tac (ZF_cs addIs [lepoll_0_is_0, lepoll_refl]) 1);
+by (fast_tac (!claset addIs [lepoll_0_is_0, lepoll_refl]) 1);
qed "lepoll_0_iff";
goalw Cardinal.thy [lepoll_def]
"!!A. [| A lepoll B; C lepoll D; B Int D = 0 |] ==> A Un C lepoll B Un D";
-by (fast_tac (ZF_cs addIs [inj_disjoint_Un]) 1);
+by (fast_tac (!claset addIs [inj_disjoint_Un]) 1);
qed "Un_lepoll_Un";
(*A eqpoll 0 ==> A=0*)
bind_thm ("eqpoll_0_is_0", eqpoll_imp_lepoll RS lepoll_0_is_0);
goal Cardinal.thy "A eqpoll 0 <-> A=0";
-by (fast_tac (ZF_cs addIs [eqpoll_0_is_0, eqpoll_refl]) 1);
+by (fast_tac (!claset addIs [eqpoll_0_is_0, eqpoll_refl]) 1);
qed "eqpoll_0_iff";
goalw Cardinal.thy [eqpoll_def]
"!!A. [| A eqpoll B; C eqpoll D; A Int C = 0; B Int D = 0 |] ==> \
\ A Un C eqpoll B Un D";
-by (fast_tac (ZF_cs addIs [bij_disjoint_Un]) 1);
+by (fast_tac (!claset addIs [bij_disjoint_Un]) 1);
qed "eqpoll_disjoint_Un";
(*** lesspoll: contributions by Krzysztof Grabczewski ***)
goalw Cardinal.thy [lesspoll_def] "!!A. A lesspoll B ==> A lepoll B";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "lesspoll_imp_lepoll";
goalw Cardinal.thy [lepoll_def]
"!!A. [| A lepoll B; well_ord(B,r) |] ==> EX s. well_ord(A,s)";
-by (fast_tac (ZF_cs addSEs [well_ord_rvimage]) 1);
+by (fast_tac (!claset addSEs [well_ord_rvimage]) 1);
qed "lepoll_well_ord";
goalw Cardinal.thy [lesspoll_def] "A lepoll B <-> A lesspoll B | A eqpoll B";
-by (fast_tac (ZF_cs addSIs [eqpollI] addSEs [eqpollE]) 1);
+by (fast_tac (!claset addSIs [eqpollI] addSEs [eqpollE]) 1);
qed "lepoll_iff_leqpoll";
goalw Cardinal.thy [inj_def, surj_def]
"!!f. [| f : inj(A, succ(m)); f ~: surj(A, succ(m)) |] ==> EX f. f:inj(A,m)";
-by (safe_tac lemmas_cs);
+by (safe_tac (claset_of"ZF"));
by (swap_res_tac [exI] 1);
by (res_inst_tac [("a", "lam z:A. if(f`z=m, y, f`z)")] CollectI 1);
-by (fast_tac (ZF_cs addSIs [if_type RS lam_type]
- addEs [apply_funtype RS succE]) 1);
+by (fast_tac (!claset addSIs [if_type RS lam_type]
+ addEs [apply_funtype RS succE]) 1);
(*Proving it's injective*)
-by (asm_simp_tac (ZF_ss setloop split_tac [expand_if]) 1);
-by (fast_tac ZF_cs 1);
+by (asm_simp_tac (!simpset setloop split_tac [expand_if]) 1);
+by (Fast_tac 1);
qed "inj_not_surj_succ";
(** Variations on transitivity **)
goalw Cardinal.thy [lesspoll_def]
"!!X. [| X lesspoll Y; Y lesspoll Z |] ==> X lesspoll Z";
-by (fast_tac (ZF_cs addSIs [eqpollI] addSEs [eqpollE] addIs [lepoll_trans]) 1);
+by (fast_tac (!claset addSIs [eqpollI] addSEs [eqpollE] addIs [lepoll_trans]) 1);
qed "lesspoll_trans";
goalw Cardinal.thy [lesspoll_def]
"!!X. [| X lesspoll Y; Y lepoll Z |] ==> X lesspoll Z";
-by (fast_tac (ZF_cs addSIs [eqpollI] addSEs [eqpollE] addIs [lepoll_trans]) 1);
+by (fast_tac (!claset addSIs [eqpollI] addSEs [eqpollE] addIs [lepoll_trans]) 1);
qed "lesspoll_lepoll_lesspoll";
goalw Cardinal.thy [lesspoll_def]
"!!X. [| X lesspoll Y; Z lepoll X |] ==> Z lesspoll Y";
-by (fast_tac (ZF_cs addSIs [eqpollI] addSEs [eqpollE] addIs [lepoll_trans]) 1);
+by (fast_tac (!claset addSIs [eqpollI] addSEs [eqpollE] addIs [lepoll_trans]) 1);
qed "lepoll_lesspoll_lesspoll";
@@ -189,10 +189,10 @@
val [premP,premOrd,premNot] = goalw Cardinal.thy [Least_def]
"[| P(i); Ord(i); !!x. x<i ==> ~P(x) |] ==> (LEAST x.P(x)) = i";
by (rtac the_equality 1);
-by (fast_tac (ZF_cs addSIs [premP,premOrd,premNot]) 1);
+by (fast_tac (!claset addSIs [premP,premOrd,premNot]) 1);
by (REPEAT (etac conjE 1));
by (etac (premOrd RS Ord_linear_lt) 1);
-by (ALLGOALS (fast_tac (ZF_cs addSIs [premP] addSDs [premNot])));
+by (ALLGOALS (fast_tac (!claset addSIs [premP] addSDs [premNot])));
qed "Least_equality";
goal Cardinal.thy "!!i. [| P(i); Ord(i) |] ==> P(LEAST x.P(x))";
@@ -202,7 +202,7 @@
by (rtac classical 1);
by (EVERY1 [stac Least_equality, assume_tac, assume_tac]);
by (assume_tac 2);
-by (fast_tac (ZF_cs addSEs [ltE]) 1);
+by (fast_tac (!claset addSEs [ltE]) 1);
qed "LeastI";
(*Proof is almost identical to the one above!*)
@@ -213,7 +213,7 @@
by (rtac classical 1);
by (EVERY1 [stac Least_equality, assume_tac, assume_tac]);
by (etac le_refl 2);
-by (fast_tac (ZF_cs addEs [ltE, lt_trans1] addIs [leI, ltI]) 1);
+by (fast_tac (!claset addEs [ltE, lt_trans1] addIs [leI, ltI]) 1);
qed "Least_le";
(*LEAST really is the smallest*)
@@ -234,12 +234,12 @@
goalw Cardinal.thy [Least_def]
"!!P. [| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x.P(x)) = 0";
by (rtac the_0 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Least_0";
goal Cardinal.thy "Ord(LEAST x.P(x))";
by (excluded_middle_tac "EX i. Ord(i) & P(i)" 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac (Least_le RS ltE) 2);
by (REPEAT_SOME assume_tac);
by (etac (Least_0 RS ssubst) 1);
@@ -252,14 +252,14 @@
(*Not needed for simplification, but helpful below*)
val prems = goal Cardinal.thy
"[| !!y. P(y) <-> Q(y) |] ==> (LEAST x.P(x)) = (LEAST x.Q(x))";
-by (simp_tac (FOL_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
qed "Least_cong";
(*Need AC to get X lepoll Y ==> |X| le |Y|; see well_ord_lepoll_imp_Card_le
Converse also requires AC, but see well_ord_cardinal_eqE*)
goalw Cardinal.thy [eqpoll_def,cardinal_def] "!!X Y. X eqpoll Y ==> |X| = |Y|";
by (rtac Least_cong 1);
-by (fast_tac (ZF_cs addEs [comp_bij,bij_converse_bij]) 1);
+by (fast_tac (!claset addEs [comp_bij,bij_converse_bij]) 1);
qed "cardinal_cong";
(*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*)
@@ -277,12 +277,12 @@
"!!X Y. [| well_ord(X,r); well_ord(Y,s); |X| = |Y| |] ==> X eqpoll Y";
by (rtac (eqpoll_sym RS eqpoll_trans) 1);
by (etac well_ord_cardinal_eqpoll 1);
-by (asm_simp_tac (ZF_ss addsimps [well_ord_cardinal_eqpoll]) 1);
+by (asm_simp_tac (!simpset addsimps [well_ord_cardinal_eqpoll]) 1);
qed "well_ord_cardinal_eqE";
goal Cardinal.thy
"!!X Y. [| well_ord(X,r); well_ord(Y,s) |] ==> |X| = |Y| <-> X eqpoll Y";
-by (fast_tac (ZF_cs addSEs [cardinal_cong, well_ord_cardinal_eqE]) 1);
+by (fast_tac (!claset addSEs [cardinal_cong, well_ord_cardinal_eqE]) 1);
qed "well_ord_cardinal_eqpoll_iff";
@@ -309,7 +309,7 @@
qed "Card_is_Ord";
goal Cardinal.thy "!!K. Card(K) ==> K le |K|";
-by (asm_simp_tac (ZF_ss addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
+by (asm_simp_tac (!simpset addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
qed "Card_cardinal_le";
goalw Cardinal.thy [cardinal_def] "Ord(|A|)";
@@ -318,8 +318,8 @@
(*The cardinals are the initial ordinals*)
goal Cardinal.thy "Card(K) <-> Ord(K) & (ALL j. j<K --> ~ j eqpoll K)";
-by (safe_tac (ZF_cs addSIs [CardI, Card_is_Ord]));
-by (fast_tac ZF_cs 2);
+by (safe_tac (!claset addSIs [CardI, Card_is_Ord]));
+by (Fast_tac 2);
by (rewrite_goals_tac [Card_def, cardinal_def]);
by (rtac less_LeastE 1);
by (etac subst 2);
@@ -328,21 +328,21 @@
goalw Cardinal.thy [lesspoll_def] "!!a. [| Card(a); i<a |] ==> i lesspoll a";
by (dresolve_tac [Card_iff_initial RS iffD1] 1);
-by (fast_tac (ZF_cs addSEs [leI RS le_imp_lepoll]) 1);
+by (fast_tac (!claset addSEs [leI RS le_imp_lepoll]) 1);
qed "lt_Card_imp_lesspoll";
goal Cardinal.thy "Card(0)";
by (rtac (Ord_0 RS CardI) 1);
-by (fast_tac (ZF_cs addSEs [ltE]) 1);
+by (fast_tac (!claset addSEs [ltE]) 1);
qed "Card_0";
val [premK,premL] = goal Cardinal.thy
"[| Card(K); Card(L) |] ==> Card(K Un L)";
by (rtac ([premK RS Card_is_Ord, premL RS Card_is_Ord] MRS Ord_linear_le) 1);
by (asm_simp_tac
- (ZF_ss addsimps [premL, le_imp_subset, subset_Un_iff RS iffD1]) 1);
+ (!simpset addsimps [premL, le_imp_subset, subset_Un_iff RS iffD1]) 1);
by (asm_simp_tac
- (ZF_ss addsimps [premK, le_imp_subset, subset_Un_iff2 RS iffD1]) 1);
+ (!simpset addsimps [premK, le_imp_subset, subset_Un_iff2 RS iffD1]) 1);
qed "Card_Un";
(*Infinite unions of cardinals? See Devlin, Lemma 6.7, page 98*)
@@ -351,7 +351,7 @@
by (excluded_middle_tac "EX i. Ord(i) & i eqpoll A" 1);
by (etac (Least_0 RS ssubst) 1 THEN rtac Card_0 1);
by (rtac (Ord_Least RS CardI) 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac less_LeastE 1);
by (assume_tac 2);
by (etac eqpoll_trans 1);
@@ -388,16 +388,16 @@
qed "cardinal_lt_imp_lt";
goal Cardinal.thy "!!i j. [| |i| < K; Ord(i); Card(K) |] ==> i < K";
-by (asm_simp_tac (ZF_ss addsimps
+by (asm_simp_tac (!simpset addsimps
[cardinal_lt_imp_lt, Card_is_Ord, Card_cardinal_eq]) 1);
qed "Card_lt_imp_lt";
goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (|i| < K) <-> (i < K)";
-by (fast_tac (ZF_cs addEs [Card_lt_imp_lt, Ord_cardinal_le RS lt_trans1]) 1);
+by (fast_tac (!claset addEs [Card_lt_imp_lt, Ord_cardinal_le RS lt_trans1]) 1);
qed "Card_lt_iff";
goal Cardinal.thy "!!i j. [| Ord(i); Card(K) |] ==> (K le |i|) <-> (K le i)";
-by (asm_simp_tac (ZF_ss addsimps
+by (asm_simp_tac (!simpset addsimps
[Card_lt_iff, Card_is_Ord, Ord_cardinal,
not_lt_iff_le RS iff_sym]) 1);
qed "Card_le_iff";
@@ -433,22 +433,22 @@
goalw Cardinal.thy [lepoll_def, inj_def]
"!!A B. [| cons(u,A) lepoll cons(v,B); u~:A; v~:B |] ==> A lepoll B";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (res_inst_tac [("x", "lam x:A. if(f`x=v, f`u, f`x)")] exI 1);
by (rtac CollectI 1);
(*Proving it's in the function space A->B*)
by (rtac (if_type RS lam_type) 1);
-by (fast_tac (ZF_cs addEs [apply_funtype RS consE]) 1);
-by (fast_tac (ZF_cs addSEs [mem_irrefl] addEs [apply_funtype RS consE]) 1);
+by (fast_tac (!claset addEs [apply_funtype RS consE]) 1);
+by (fast_tac (!claset addSEs [mem_irrefl] addEs [apply_funtype RS consE]) 1);
(*Proving it's injective*)
-by (asm_simp_tac (ZF_ss setloop split_tac [expand_if]) 1);
-by (fast_tac ZF_cs 1);
+by (asm_simp_tac (!simpset setloop split_tac [expand_if]) 1);
+by (Fast_tac 1);
qed "cons_lepoll_consD";
goal Cardinal.thy
"!!A B. [| cons(u,A) eqpoll cons(v,B); u~:A; v~:B |] ==> A eqpoll B";
-by (asm_full_simp_tac (ZF_ss addsimps [eqpoll_iff]) 1);
-by (fast_tac (ZF_cs addIs [cons_lepoll_consD]) 1);
+by (asm_full_simp_tac (!simpset addsimps [eqpoll_iff]) 1);
+by (fast_tac (!claset addIs [cons_lepoll_consD]) 1);
qed "cons_eqpoll_consD";
(*Lemma suggested by Mike Fourman*)
@@ -460,12 +460,12 @@
val [prem] = goal Cardinal.thy
"m:nat ==> ALL n: nat. m lepoll n --> m le n";
by (nat_ind_tac "m" [prem] 1);
-by (fast_tac (ZF_cs addSIs [nat_0_le]) 1);
+by (fast_tac (!claset addSIs [nat_0_le]) 1);
by (rtac ballI 1);
by (eres_inst_tac [("n","n")] natE 1);
-by (asm_simp_tac (ZF_ss addsimps [lepoll_def, inj_def,
+by (asm_simp_tac (!simpset addsimps [lepoll_def, inj_def,
succI1 RS Pi_empty2]) 1);
-by (fast_tac (ZF_cs addSIs [succ_leI] addSDs [succ_lepoll_succD]) 1);
+by (fast_tac (!claset addSIs [succ_leI] addSDs [succ_lepoll_succD]) 1);
qed "nat_lepoll_imp_le_lemma";
bind_thm ("nat_lepoll_imp_le", nat_lepoll_imp_le_lemma RS bspec RS mp);
@@ -473,8 +473,8 @@
goal Cardinal.thy
"!!m n. [| m:nat; n: nat |] ==> m eqpoll n <-> m = n";
by (rtac iffI 1);
-by (asm_simp_tac (ZF_ss addsimps [eqpoll_refl]) 2);
-by (fast_tac (ZF_cs addIs [nat_lepoll_imp_le, le_anti_sym]
+by (asm_simp_tac (!simpset addsimps [eqpoll_refl]) 2);
+by (fast_tac (!claset addIs [nat_lepoll_imp_le, le_anti_sym]
addSEs [eqpollE]) 1);
qed "nat_eqpoll_iff";
@@ -483,8 +483,8 @@
"!!n. n: nat ==> Card(n)";
by (stac Least_equality 1);
by (REPEAT_FIRST (ares_tac [eqpoll_refl, nat_into_Ord, refl]));
-by (asm_simp_tac (ZF_ss addsimps [lt_nat_in_nat RS nat_eqpoll_iff]) 1);
-by (fast_tac (ZF_cs addSEs [lt_irrefl]) 1);
+by (asm_simp_tac (!simpset addsimps [lt_nat_in_nat RS nat_eqpoll_iff]) 1);
+by (fast_tac (!claset addSEs [lt_irrefl]) 1);
qed "nat_into_Card";
(*Part of Kunen's Lemma 10.6*)
@@ -499,7 +499,7 @@
goalw Cardinal.thy [lesspoll_def]
"!!m. [| A lepoll m; m:nat |] ==> A lesspoll succ(m)";
by (rtac conjI 1);
-by (fast_tac (ZF_cs addIs [subset_imp_lepoll RSN (2,lepoll_trans)]) 1);
+by (fast_tac (!claset addIs [subset_imp_lepoll RSN (2,lepoll_trans)]) 1);
by (rtac notI 1);
by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll] 1);
by (dtac lepoll_trans 1 THEN assume_tac 1);
@@ -508,12 +508,12 @@
goalw Cardinal.thy [lesspoll_def, lepoll_def, eqpoll_def, bij_def]
"!!m. [| A lesspoll succ(m); m:nat |] ==> A lepoll m";
-by (step_tac ZF_cs 1);
-by (fast_tac (ZF_cs addSIs [inj_not_surj_succ]) 1);
+by (step_tac (!claset) 1);
+by (fast_tac (!claset addSIs [inj_not_surj_succ]) 1);
qed "lesspoll_succ_imp_lepoll";
goal Cardinal.thy "!!m. m:nat ==> A lesspoll succ(m) <-> A lepoll m";
-by (fast_tac (ZF_cs addSIs [lepoll_imp_lesspoll_succ,
+by (fast_tac (!claset addSIs [lepoll_imp_lesspoll_succ,
lesspoll_succ_imp_lepoll]) 1);
qed "lesspoll_succ_iff";
@@ -522,7 +522,7 @@
by (rtac disjCI 1);
by (rtac lesspoll_succ_imp_lepoll 1);
by (assume_tac 2);
-by (asm_simp_tac (ZF_ss addsimps [lesspoll_def]) 1);
+by (asm_simp_tac (!simpset addsimps [lesspoll_def]) 1);
qed "lepoll_succ_disj";
@@ -539,7 +539,7 @@
goal Cardinal.thy "!!i n. [| Ord(i); n:nat |] ==> i eqpoll n <-> i=n";
by (rtac iffI 1);
-by (asm_simp_tac (ZF_ss addsimps [eqpoll_refl]) 2);
+by (asm_simp_tac (!simpset addsimps [eqpoll_refl]) 2);
by (rtac Ord_linear_lt 1);
by (REPEAT_SOME (eresolve_tac [asm_rl, nat_into_Ord]));
by (etac (lt_nat_in_nat RS nat_eqpoll_iff RS iffD1) 1 THEN
@@ -552,7 +552,7 @@
by (stac Least_equality 1);
by (REPEAT_FIRST (ares_tac [eqpoll_refl, Ord_nat, refl]));
by (etac ltE 1);
-by (asm_simp_tac (ZF_ss addsimps [eqpoll_iff, lt_not_lepoll, ltI]) 1);
+by (asm_simp_tac (!simpset addsimps [eqpoll_iff, lt_not_lepoll, ltI]) 1);
qed "Card_nat";
(*Allows showing that |i| is a limit cardinal*)
@@ -568,40 +568,40 @@
(*Congruence law for cons under equipollence*)
goalw Cardinal.thy [lepoll_def]
"!!A B. [| A lepoll B; b ~: B |] ==> cons(a,A) lepoll cons(b,B)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (res_inst_tac [("x", "lam y: cons(a,A).if(y=a, b, f`y)")] exI 1);
by (res_inst_tac [("d","%z.if(z:B, converse(f)`z, a)")]
lam_injective 1);
-by (asm_simp_tac (ZF_ss addsimps [inj_is_fun RS apply_type, cons_iff]
+by (asm_simp_tac (!simpset addsimps [inj_is_fun RS apply_type, cons_iff]
setloop etac consE') 1);
-by (asm_simp_tac (ZF_ss addsimps [inj_is_fun RS apply_type, left_inverse]
+by (asm_simp_tac (!simpset addsimps [inj_is_fun RS apply_type, left_inverse]
setloop etac consE') 1);
qed "cons_lepoll_cong";
goal Cardinal.thy
"!!A B. [| A eqpoll B; a ~: A; b ~: B |] ==> cons(a,A) eqpoll cons(b,B)";
-by (asm_full_simp_tac (ZF_ss addsimps [eqpoll_iff, cons_lepoll_cong]) 1);
+by (asm_full_simp_tac (!simpset addsimps [eqpoll_iff, cons_lepoll_cong]) 1);
qed "cons_eqpoll_cong";
goal Cardinal.thy
"!!A B. [| a ~: A; b ~: B |] ==> \
\ cons(a,A) lepoll cons(b,B) <-> A lepoll B";
-by (fast_tac (ZF_cs addIs [cons_lepoll_cong, cons_lepoll_consD]) 1);
+by (fast_tac (!claset addIs [cons_lepoll_cong, cons_lepoll_consD]) 1);
qed "cons_lepoll_cons_iff";
goal Cardinal.thy
"!!A B. [| a ~: A; b ~: B |] ==> \
\ cons(a,A) eqpoll cons(b,B) <-> A eqpoll B";
-by (fast_tac (ZF_cs addIs [cons_eqpoll_cong, cons_eqpoll_consD]) 1);
+by (fast_tac (!claset addIs [cons_eqpoll_cong, cons_eqpoll_consD]) 1);
qed "cons_eqpoll_cons_iff";
goalw Cardinal.thy [succ_def] "{a} eqpoll 1";
-by (fast_tac (ZF_cs addSIs [eqpoll_refl RS cons_eqpoll_cong]) 1);
+by (fast_tac (!claset addSIs [eqpoll_refl RS cons_eqpoll_cong]) 1);
qed "singleton_eqpoll_1";
goal Cardinal.thy "|{a}| = 1";
by (resolve_tac [singleton_eqpoll_1 RS cardinal_cong RS trans] 1);
-by (simp_tac (arith_ss addsimps [nat_into_Card RS Card_cardinal_eq]) 1);
+by (simp_tac (!simpset addsimps [nat_into_Card RS Card_cardinal_eq]) 1);
qed "cardinal_singleton";
(*Congruence law for succ under equipollence*)
@@ -613,13 +613,13 @@
(*Congruence law for + under equipollence*)
goalw Cardinal.thy [eqpoll_def]
"!!A B C D. [| A eqpoll C; B eqpoll D |] ==> A+B eqpoll C+D";
-by (fast_tac (ZF_cs addSIs [sum_bij]) 1);
+by (fast_tac (!claset addSIs [sum_bij]) 1);
qed "sum_eqpoll_cong";
(*Congruence law for * under equipollence*)
goalw Cardinal.thy [eqpoll_def]
"!!A B C D. [| A eqpoll C; B eqpoll D |] ==> A*B eqpoll C*D";
-by (fast_tac (ZF_cs addSIs [prod_bij]) 1);
+by (fast_tac (!claset addSIs [prod_bij]) 1);
qed "prod_eqpoll_cong";
goalw Cardinal.thy [eqpoll_def]
@@ -628,16 +628,16 @@
by (res_inst_tac [("c", "%x. if(x:A, f`x, x)"),
("d", "%y. if(y: range(f), converse(f)`y, y)")]
lam_bijective 1);
-by (fast_tac (ZF_cs addSIs [if_type, apply_type] addIs [inj_is_fun]) 1);
+by (fast_tac (!claset addSIs [if_type, apply_type] addIs [inj_is_fun]) 1);
by (asm_simp_tac
- (ZF_ss addsimps [inj_converse_fun RS apply_funtype]
+ (!simpset addsimps [inj_converse_fun RS apply_funtype]
setloop split_tac [expand_if]) 1);
-by (asm_simp_tac (ZF_ss addsimps [inj_is_fun RS apply_rangeI, left_inverse]
+by (asm_simp_tac (!simpset addsimps [inj_is_fun RS apply_rangeI, left_inverse]
setloop etac UnE') 1);
by (asm_simp_tac
- (ZF_ss addsimps [inj_converse_fun RS apply_funtype, right_inverse]
+ (!simpset addsimps [inj_converse_fun RS apply_funtype, right_inverse]
setloop split_tac [expand_if]) 1);
-by (fast_tac (ZF_cs addEs [equals0D]) 1);
+by (fast_tac (!claset addEs [equals0D]) 1);
qed "inj_disjoint_eqpoll";
@@ -650,7 +650,7 @@
by (rtac cons_lepoll_consD 1);
by (rtac mem_not_refl 3);
by (eresolve_tac [cons_Diff RS ssubst] 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
qed "Diff_sing_lepoll";
(*If A has at least n+1 elements then A-{a} has at least n.*)
@@ -658,12 +658,12 @@
"!!A a n. [| succ(n) lepoll A |] ==> n lepoll A - {a}";
by (rtac cons_lepoll_consD 1);
by (rtac mem_not_refl 2);
-by (fast_tac ZF_cs 2);
-by (fast_tac (ZF_cs addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
+by (Fast_tac 2);
+by (fast_tac (!claset addSEs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
qed "lepoll_Diff_sing";
goal Cardinal.thy "!!A a n. [| a:A; A eqpoll succ(n) |] ==> A - {a} eqpoll n";
-by (fast_tac (ZF_cs addSIs [eqpollI] addSEs [eqpollE]
+by (fast_tac (!claset addSIs [eqpollI] addSEs [eqpollE]
addIs [Diff_sing_lepoll,lepoll_Diff_sing]) 1);
qed "Diff_sing_eqpoll";
@@ -678,8 +678,8 @@
by (res_inst_tac [("x","lam x: A Un B. if (x:A,Inl(x),Inr(x))")] exI 1);
by (res_inst_tac [("d","%z. snd(z)")] lam_injective 1);
by (split_tac [expand_if] 1);
-by (fast_tac (ZF_cs addSIs [InlI, InrI]) 1);
-by (asm_full_simp_tac (ZF_ss addsimps [Inl_def, Inr_def]
+by (fast_tac (!claset addSIs [InlI, InrI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [Inl_def, Inr_def]
setloop split_tac [expand_if]) 1);
qed "Un_lepoll_sum";
@@ -687,20 +687,20 @@
(*** Finite and infinite sets ***)
goalw Cardinal.thy [Finite_def] "Finite(0)";
-by (fast_tac (ZF_cs addSIs [eqpoll_refl, nat_0I]) 1);
+by (fast_tac (!claset addSIs [eqpoll_refl, nat_0I]) 1);
qed "Finite_0";
goalw Cardinal.thy [Finite_def]
"!!A. [| A lepoll n; n:nat |] ==> Finite(A)";
by (etac rev_mp 1);
by (etac nat_induct 1);
-by (fast_tac (ZF_cs addSDs [lepoll_0_is_0] addSIs [eqpoll_refl,nat_0I]) 1);
-by (fast_tac (ZF_cs addSDs [lepoll_succ_disj] addSIs [nat_succI]) 1);
+by (fast_tac (!claset addSDs [lepoll_0_is_0] addSIs [eqpoll_refl,nat_0I]) 1);
+by (fast_tac (!claset addSDs [lepoll_succ_disj] addSIs [nat_succI]) 1);
qed "lepoll_nat_imp_Finite";
goalw Cardinal.thy [Finite_def]
"!!X. [| Y lepoll X; Finite(X) |] ==> Finite(Y)";
-by (fast_tac (ZF_cs addSEs [eqpollE]
+by (fast_tac (!claset addSEs [eqpollE]
addEs [lepoll_trans RS
rewrite_rule [Finite_def] lepoll_nat_imp_Finite]) 1);
qed "lepoll_Finite";
@@ -711,12 +711,12 @@
goalw Cardinal.thy [Finite_def] "!!x. Finite(x) ==> Finite(cons(y,x))";
by (excluded_middle_tac "y:x" 1);
-by (asm_simp_tac (ZF_ss addsimps [cons_absorb]) 2);
+by (asm_simp_tac (!simpset addsimps [cons_absorb]) 2);
by (etac bexE 1);
by (rtac bexI 1);
by (etac nat_succI 2);
by (asm_simp_tac
- (ZF_ss addsimps [succ_def, cons_eqpoll_cong, mem_not_refl]) 1);
+ (!simpset addsimps [succ_def, cons_eqpoll_cong, mem_not_refl]) 1);
qed "Finite_cons";
goalw Cardinal.thy [succ_def] "!!x. Finite(x) ==> Finite(succ(x))";
@@ -727,12 +727,12 @@
"!!i. [| Ord(i); ~ Finite(i) |] ==> nat le i";
by (eresolve_tac [Ord_nat RSN (2,Ord_linear2)] 1);
by (assume_tac 2);
-by (fast_tac (ZF_cs addSIs [eqpoll_refl] addSEs [ltE]) 1);
+by (fast_tac (!claset addSIs [eqpoll_refl] addSEs [ltE]) 1);
qed "nat_le_infinite_Ord";
goalw Cardinal.thy [Finite_def, eqpoll_def]
"!!A. Finite(A) ==> EX r. well_ord(A,r)";
-by (fast_tac (ZF_cs addIs [well_ord_rvimage, bij_is_inj, well_ord_Memrel,
+by (fast_tac (!claset addIs [well_ord_rvimage, bij_is_inj, well_ord_Memrel,
nat_into_Ord]) 1);
qed "Finite_imp_well_ord";
@@ -742,22 +742,22 @@
goal Nat.thy "!!n. n:nat ==> wf[n](converse(Memrel(n)))";
by (etac nat_induct 1);
-by (fast_tac (ZF_cs addIs [wf_onI]) 1);
+by (fast_tac (!claset addIs [wf_onI]) 1);
by (rtac wf_onI 1);
-by (asm_full_simp_tac (ZF_ss addsimps [wf_on_def, wf_def, Memrel_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [wf_on_def, wf_def, Memrel_iff]) 1);
by (excluded_middle_tac "x:Z" 1);
by (dres_inst_tac [("x", "x")] bspec 2 THEN assume_tac 2);
-by (fast_tac (ZF_cs addSEs [mem_irrefl] addEs [mem_asym]) 2);
+by (fast_tac (!claset addSEs [mem_irrefl] addEs [mem_asym]) 2);
by (dres_inst_tac [("x", "Z")] spec 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (dres_inst_tac [("x", "xa")] bspec 1 THEN assume_tac 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "nat_wf_on_converse_Memrel";
goal Cardinal.thy "!!n. n:nat ==> well_ord(n,converse(Memrel(n)))";
by (forward_tac [Ord_nat RS Ord_in_Ord RS well_ord_Memrel] 1);
by (rewtac well_ord_def);
-by (fast_tac (ZF_cs addSIs [tot_ord_converse, nat_wf_on_converse_Memrel]) 1);
+by (fast_tac (!claset addSIs [tot_ord_converse, nat_wf_on_converse_Memrel]) 1);
qed "nat_well_ord_converse_Memrel";
goal Cardinal.thy
@@ -768,7 +768,7 @@
by (forward_tac [ordermap_bij RS bij_is_inj RS well_ord_rvimage] 1);
by (assume_tac 1);
by (asm_full_simp_tac
- (ZF_ss addsimps [rvimage_converse, converse_Int, converse_prod,
+ (!simpset addsimps [rvimage_converse, converse_Int, converse_prod,
ordertype_ord_iso RS ord_iso_rvimage_eq]) 1);
qed "well_ord_converse";
@@ -778,12 +778,12 @@
REPEAT (assume_tac 1));
by (rtac eqpoll_trans 1 THEN assume_tac 2);
by (rewtac eqpoll_def);
-by (fast_tac (ZF_cs addSIs [ordermap_bij RS bij_converse_bij]) 1);
+by (fast_tac (!claset addSIs [ordermap_bij RS bij_converse_bij]) 1);
qed "ordertype_eq_n";
goalw Cardinal.thy [Finite_def]
"!!A. [| Finite(A); well_ord(A,r) |] ==> well_ord(A,converse(r))";
by (rtac well_ord_converse 1 THEN assume_tac 1);
-by (fast_tac (ZF_cs addDs [ordertype_eq_n]
+by (fast_tac (!claset addDs [ordertype_eq_n]
addSIs [nat_well_ord_converse_Memrel]) 1);
qed "Finite_well_ord_converse";
--- a/src/ZF/CardinalArith.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/CardinalArith.ML Fri Jan 03 15:01:55 1997 +0100
@@ -21,8 +21,8 @@
by (rtac exI 1);
by (res_inst_tac [("c", "case(Inr, Inl)"), ("d", "case(Inr, Inl)")]
lam_bijective 1);
-by (safe_tac (ZF_cs addSEs [sumE]));
-by (ALLGOALS (asm_simp_tac case_ss));
+by (safe_tac (!claset addSEs [sumE]));
+by (ALLGOALS (Asm_simp_tac));
qed "sum_commute_eqpoll";
goalw CardinalArith.thy [cadd_def] "i |+| j = j |+| i";
@@ -57,7 +57,7 @@
qed "sum_0_eqpoll";
goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
-by (asm_simp_tac (ZF_ss addsimps [sum_0_eqpoll RS cardinal_cong,
+by (asm_simp_tac (!simpset addsimps [sum_0_eqpoll RS cardinal_cong,
Card_cardinal_eq]) 1);
qed "cadd_0";
@@ -65,7 +65,7 @@
goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A+B";
by (res_inst_tac [("x", "lam x:A. Inl(x)")] exI 1);
-by (asm_simp_tac (sum_ss addsimps [lam_type]) 1);
+by (asm_simp_tac (!simpset addsimps [lam_type]) 1);
qed "sum_lepoll_self";
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
@@ -88,12 +88,12 @@
lam_injective 1);
by (typechk_tac ([inj_is_fun, case_type, InlI, InrI] @ ZF_typechecks));
by (etac sumE 1);
-by (ALLGOALS (asm_simp_tac (sum_ss addsimps [left_inverse])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [left_inverse])));
qed "sum_lepoll_mono";
goalw CardinalArith.thy [cadd_def]
"!!K. [| K' le K; L' le L |] ==> (K' |+| L') le (K |+| L)";
-by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1]));
+by (safe_tac (!claset addSDs [le_subset_iff RS iffD1]));
by (rtac well_ord_lepoll_imp_Card_le 1);
by (REPEAT (ares_tac [sum_lepoll_mono, subset_imp_lepoll] 2));
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1));
@@ -107,7 +107,7 @@
("d", "%z.if(z=A+B,Inl(A),z)")]
lam_bijective 1);
by (ALLGOALS
- (asm_simp_tac (case_ss addsimps [succI2, mem_imp_not_eq]
+ (asm_simp_tac (!simpset addsimps [succI2, mem_imp_not_eq]
setloop eresolve_tac [sumE,succE])));
qed "sum_succ_eqpoll";
@@ -125,8 +125,8 @@
"[| m: nat; n: nat |] ==> m |+| n = m#+n";
by (cut_facts_tac [nnat] 1);
by (nat_ind_tac "m" [mnat] 1);
-by (asm_simp_tac (arith_ss addsimps [nat_into_Card RS cadd_0]) 1);
-by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cadd_succ_lemma,
+by (asm_simp_tac (!simpset addsimps [nat_into_Card RS cadd_0]) 1);
+by (asm_simp_tac (!simpset addsimps [nat_into_Ord, cadd_succ_lemma,
nat_into_Card RS Card_cardinal_eq]) 1);
qed "nat_cadd_eq_add";
@@ -140,8 +140,8 @@
by (rtac exI 1);
by (res_inst_tac [("c", "%<x,y>.<y,x>"), ("d", "%<x,y>.<y,x>")]
lam_bijective 1);
-by (safe_tac ZF_cs);
-by (ALLGOALS (asm_simp_tac ZF_ss));
+by (safe_tac (!claset));
+by (ALLGOALS (Asm_simp_tac));
qed "prod_commute_eqpoll";
goalw CardinalArith.thy [cmult_def] "i |*| j = j |*| i";
@@ -192,11 +192,11 @@
goalw CardinalArith.thy [eqpoll_def] "0*A eqpoll 0";
by (rtac exI 1);
by (rtac lam_bijective 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
qed "prod_0_eqpoll";
goalw CardinalArith.thy [cmult_def] "0 |*| i = 0";
-by (asm_simp_tac (ZF_ss addsimps [prod_0_eqpoll RS cardinal_cong,
+by (asm_simp_tac (!simpset addsimps [prod_0_eqpoll RS cardinal_cong,
Card_0 RS Card_cardinal_eq]) 1);
qed "cmult_0";
@@ -208,7 +208,7 @@
qed "prod_singleton_eqpoll";
goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K";
-by (asm_simp_tac (ZF_ss addsimps [prod_singleton_eqpoll RS cardinal_cong,
+by (asm_simp_tac (!simpset addsimps [prod_singleton_eqpoll RS cardinal_cong,
Card_cardinal_eq]) 1);
qed "cmult_1";
@@ -216,7 +216,7 @@
goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A*A";
by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1);
-by (simp_tac (ZF_ss addsimps [lam_type]) 1);
+by (simp_tac (!simpset addsimps [lam_type]) 1);
qed "prod_square_lepoll";
(*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
@@ -225,14 +225,14 @@
by (rtac well_ord_lepoll_imp_Card_le 2);
by (rtac prod_square_lepoll 3);
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord] 2));
-by (asm_simp_tac (ZF_ss addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
+by (asm_simp_tac (!simpset addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
qed "cmult_square_le";
(** Multiplication by a non-zero cardinal **)
goalw CardinalArith.thy [lepoll_def, inj_def] "!!b. b: B ==> A lepoll A*B";
by (res_inst_tac [("x", "lam x:A. <x,b>")] exI 1);
-by (asm_simp_tac (ZF_ss addsimps [lam_type]) 1);
+by (asm_simp_tac (!simpset addsimps [lam_type]) 1);
qed "prod_lepoll_self";
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
@@ -253,12 +253,12 @@
lam_injective 1);
by (typechk_tac (inj_is_fun::ZF_typechecks));
by (etac SigmaE 1);
-by (asm_simp_tac (ZF_ss addsimps [left_inverse]) 1);
+by (asm_simp_tac (!simpset addsimps [left_inverse]) 1);
qed "prod_lepoll_mono";
goalw CardinalArith.thy [cmult_def]
"!!K. [| K' le K; L' le L |] ==> (K' |*| L') le (K |*| L)";
-by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1]));
+by (safe_tac (!claset addSDs [le_subset_iff RS iffD1]));
by (rtac well_ord_lepoll_imp_Card_le 1);
by (REPEAT (ares_tac [prod_lepoll_mono, subset_imp_lepoll] 2));
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));
@@ -271,9 +271,9 @@
by (res_inst_tac [("c", "%<x,y>. if(x=A, Inl(y), Inr(<x,y>))"),
("d", "case(%y. <A,y>, %z.z)")]
lam_bijective 1);
-by (safe_tac (ZF_cs addSEs [sumE]));
+by (safe_tac (!claset addSEs [sumE]));
by (ALLGOALS
- (asm_simp_tac (case_ss addsimps [succI2, if_type, mem_imp_not_eq])));
+ (asm_simp_tac (!simpset addsimps [succI2, if_type, mem_imp_not_eq])));
qed "prod_succ_eqpoll";
(*Unconditional version requires AC*)
@@ -289,14 +289,14 @@
"[| m: nat; n: nat |] ==> m |*| n = m#*n";
by (cut_facts_tac [nnat] 1);
by (nat_ind_tac "m" [mnat] 1);
-by (asm_simp_tac (arith_ss addsimps [cmult_0]) 1);
-by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cmult_succ_lemma,
+by (asm_simp_tac (!simpset addsimps [cmult_0]) 1);
+by (asm_simp_tac (!simpset addsimps [nat_into_Ord, cmult_succ_lemma,
nat_cadd_eq_add]) 1);
qed "nat_cmult_eq_mult";
goal CardinalArith.thy "!!m n. Card(n) ==> 2 |*| n = n |+| n";
by (asm_simp_tac
- (ZF_ss addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, Card_is_Ord,
+ (!simpset addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, Card_is_Ord,
read_instantiate [("j","0")] cadd_commute, cadd_0]) 1);
qed "cmult_2";
@@ -316,10 +316,10 @@
by (res_inst_tac [("d", "%y. if(y: range(f), \
\ nat_case(u, %z.f`z, converse(f)`y), y)")]
lam_injective 1);
-by (fast_tac (ZF_cs addSIs [if_type, nat_0I, nat_succI, apply_type]
+by (fast_tac (!claset addSIs [if_type, nat_0I, nat_succI, apply_type]
addIs [inj_is_fun, inj_converse_fun]) 1);
by (asm_simp_tac
- (ZF_ss addsimps [inj_is_fun RS apply_rangeI,
+ (!simpset addsimps [inj_is_fun RS apply_rangeI,
inj_converse_fun RS apply_rangeI,
inj_converse_fun RS apply_funtype,
left_inverse, right_inverse, nat_0I, nat_succI,
@@ -338,7 +338,7 @@
qed "nat_succ_eqpoll";
goalw CardinalArith.thy [InfCard_def] "InfCard(nat)";
-by (fast_tac (ZF_cs addIs [Card_nat, le_refl, Card_is_Ord]) 1);
+by (fast_tac (!claset addIs [Card_nat, le_refl, Card_is_Ord]) 1);
qed "InfCard_nat";
goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
@@ -347,7 +347,7 @@
goalw CardinalArith.thy [InfCard_def]
"!!K L. [| InfCard(K); Card(L) |] ==> InfCard(K Un L)";
-by (asm_simp_tac (ZF_ss addsimps [Card_Un, Un_upper1_le RSN (2,le_trans),
+by (asm_simp_tac (!simpset addsimps [Card_Un, Un_upper1_le RSN (2,le_trans),
Card_is_Ord]) 1);
qed "InfCard_Un";
@@ -357,7 +357,7 @@
by (forward_tac [Card_is_Ord] 1);
by (rtac (ltI RS non_succ_LimitI) 1);
by (etac ([asm_rl, nat_0I] MRS (le_imp_subset RS subsetD)) 1);
-by (safe_tac (ZF_cs addSDs [Limit_nat RS Limit_le_succD]));
+by (safe_tac (!claset addSDs [Limit_nat RS Limit_le_succD]));
by (rewtac Card_def);
by (dtac trans 1);
by (etac (le_imp_subset RS nat_succ_eqpoll RS cardinal_cong) 1);
@@ -372,7 +372,7 @@
goalw Cardinal.thy [eqpoll_def]
"!!A. [| well_ord(A,r); x:A |] ==> ordermap(A,r)`x eqpoll pred(A,x,r)";
by (rtac exI 1);
-by (asm_simp_tac (ZF_ss addsimps [ordermap_eq_image, well_ord_is_wf]) 1);
+by (asm_simp_tac (!simpset addsimps [ordermap_eq_image, well_ord_is_wf]) 1);
by (etac (ordermap_bij RS bij_is_inj RS restrict_bij RS bij_converse_bij) 1);
by (rtac pred_subset 1);
qed "ordermap_eqpoll_pred";
@@ -381,7 +381,7 @@
goalw CardinalArith.thy [inj_def]
"!!K. Ord(K) ==> (lam <x,y>:K*K. <x Un y, x, y>) : inj(K*K, K*K*K)";
-by (fast_tac (ZF_cs addss ZF_ss
+by (fast_tac (!claset addss (!simpset)
addIs [lam_type, Un_least_lt RS ltD, ltI]) 1);
qed "csquare_lam_inj";
@@ -397,22 +397,22 @@
"!!K. [| x<K; y<K; z<K |] ==> \
\ <<x,y>, <z,z>> : csquare_rel(K) --> x le z & y le z";
by (REPEAT (etac ltE 1));
-by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff,
+by (asm_simp_tac (!simpset addsimps [rvimage_iff, rmult_iff, Memrel_iff,
Un_absorb, Un_least_mem_iff, ltD]) 1);
-by (safe_tac (ZF_cs addSEs [mem_irrefl]
+by (safe_tac (!claset addSEs [mem_irrefl]
addSIs [Un_upper1_le, Un_upper2_le]));
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [lt_def, succI2, Ord_succ])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [lt_def, succI2, Ord_succ])));
val csquareD_lemma = result();
bind_thm ("csquareD", csquareD_lemma RS mp);
goalw CardinalArith.thy [pred_def]
"!!K. z<K ==> pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)";
-by (safe_tac (lemmas_cs addSEs [SigmaE])); (*avoids using succCI*)
+by (safe_tac (claset_of"ZF" addSEs [SigmaE])); (*avoids using succCI,...*)
by (rtac (csquareD RS conjE) 1);
by (rewtac lt_def);
by (assume_tac 4);
-by (ALLGOALS (fast_tac ZF_cs));
+by (ALLGOALS Fast_tac);
qed "pred_csquare_subset";
goalw CardinalArith.thy [csquare_rel_def]
@@ -420,8 +420,8 @@
by (subgoals_tac ["x<K", "y<K"] 1);
by (REPEAT (eresolve_tac [asm_rl, lt_trans] 2));
by (REPEAT (etac ltE 1));
-by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff,
- Un_absorb, Un_least_mem_iff, ltD]) 1);
+by (asm_simp_tac (!simpset addsimps [rvimage_iff, rmult_iff, Memrel_iff,
+ Un_absorb, Un_least_mem_iff, ltD]) 1);
qed "csquare_ltI";
(*Part of the traditional proof. UNUSED since it's harder to prove & apply *)
@@ -431,11 +431,11 @@
by (subgoals_tac ["x<K", "y<K"] 1);
by (REPEAT (eresolve_tac [asm_rl, lt_trans1] 2));
by (REPEAT (etac ltE 1));
-by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff,
+by (asm_simp_tac (!simpset addsimps [rvimage_iff, rmult_iff, Memrel_iff,
Un_absorb, Un_least_mem_iff, ltD]) 1);
by (REPEAT_FIRST (etac succE));
by (ALLGOALS
- (asm_simp_tac (ZF_ss addsimps [subset_Un_iff RS iff_sym,
+ (asm_simp_tac (!simpset addsimps [subset_Un_iff RS iff_sym,
subset_Un_iff2 RS iff_sym, OrdmemD])));
qed "csquare_or_eqI";
@@ -447,11 +447,11 @@
\ ordermap(K*K, csquare_rel(K)) ` <z,z>";
by (subgoals_tac ["z<K", "well_ord(K*K, csquare_rel(K))"] 1);
by (etac (Limit_is_Ord RS well_ord_csquare) 2);
-by (fast_tac (ZF_cs addSIs [Un_least_lt, Limit_has_succ]) 2);
+by (fast_tac (!claset addSIs [Un_least_lt, Limit_has_succ]) 2);
by (rtac (csquare_ltI RS ordermap_mono RS ltI) 1);
by (etac well_ord_is_wf 4);
by (ALLGOALS
- (fast_tac (ZF_cs addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap]
+ (fast_tac (!claset addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap]
addSEs [ltE])));
qed "ordermap_z_lt";
@@ -462,12 +462,12 @@
by (rtac (well_ord_rmult RS well_ord_lepoll_imp_Card_le) 1);
by (REPEAT (ares_tac [Ord_cardinal, well_ord_Memrel] 1));
by (subgoals_tac ["z<K"] 1);
-by (fast_tac (ZF_cs addSIs [Un_least_lt, Limit_has_succ]) 2);
+by (fast_tac (!claset addSIs [Un_least_lt, Limit_has_succ]) 2);
by (rtac (ordermap_z_lt RS leI RS le_imp_lepoll RS lepoll_trans) 1);
by (REPEAT_SOME assume_tac);
by (rtac (ordermap_eqpoll_pred RS eqpoll_imp_lepoll RS lepoll_trans) 1);
by (etac (Limit_is_Ord RS well_ord_csquare) 1);
-by (fast_tac (ZF_cs addIs [ltD]) 1);
+by (fast_tac (!claset addIs [ltD]) 1);
by (rtac (pred_csquare_subset RS subset_imp_lepoll RS lepoll_trans) 1 THEN
assume_tac 1);
by (REPEAT_FIRST (etac ltE));
@@ -486,8 +486,8 @@
by (rtac Card_lt_imp_lt 1);
by (etac InfCard_is_Card 3);
by (etac ltE 2 THEN assume_tac 2);
-by (asm_full_simp_tac (ZF_ss addsimps [ordertype_unfold]) 1);
-by (safe_tac (ZF_cs addSEs [ltE]));
+by (asm_full_simp_tac (!simpset addsimps [ordertype_unfold]) 1);
+by (safe_tac (!claset addSEs [ltE]));
by (subgoals_tac ["Ord(xb)", "Ord(y)"] 1);
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 2));
by (rtac (InfCard_is_Limit RS ordermap_csquare_le RS lt_trans1) 1 THEN
@@ -496,13 +496,13 @@
REPEAT (ares_tac [Ord_Un, Ord_nat] 1));
(*the finite case: xb Un y < nat *)
by (res_inst_tac [("j", "nat")] lt_trans2 1);
-by (asm_full_simp_tac (FOL_ss addsimps [InfCard_def]) 2);
+by (asm_full_simp_tac (!simpset addsimps [InfCard_def]) 2);
by (asm_full_simp_tac
- (ZF_ss addsimps [lt_def, nat_cmult_eq_mult, nat_succI, mult_type,
+ (!simpset addsimps [lt_def, nat_cmult_eq_mult, nat_succI, mult_type,
nat_into_Card RS Card_cardinal_eq, Ord_nat]) 1);
(*case nat le (xb Un y) *)
by (asm_full_simp_tac
- (ZF_ss addsimps [le_imp_subset RS nat_succ_eqpoll RS cardinal_cong,
+ (!simpset addsimps [le_imp_subset RS nat_succ_eqpoll RS cardinal_cong,
le_succ_iff, InfCard_def, Card_cardinal, Un_least_lt,
Ord_Un, ltI, nat_le_cardinal,
Ord_cardinal_le RS lt_trans1 RS ltD]) 1);
@@ -520,7 +520,7 @@
by (assume_tac 2);
by (assume_tac 2);
by (asm_simp_tac
- (ZF_ss addsimps [cmult_def, Ord_cardinal_le,
+ (!simpset addsimps [cmult_def, Ord_cardinal_le,
well_ord_csquare RS ordermap_bij RS
bij_imp_eqpoll RS cardinal_cong,
well_ord_csquare RS Ord_ordertype]) 1);
@@ -533,7 +533,7 @@
by (REPEAT (etac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1));
by (rtac well_ord_cardinal_eqE 1);
by (REPEAT (ares_tac [Ord_cardinal, well_ord_rmult, well_ord_Memrel] 1));
-by (asm_simp_tac (ZF_ss addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1);
+by (asm_simp_tac (!simpset addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1);
qed "well_ord_InfCard_square_eq";
(** Toward's Kunen's Corollary 10.13 (1) **)
@@ -544,7 +544,7 @@
REPEAT (ares_tac [cmult_le_self, InfCard_is_Card] 2));
by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1);
by (resolve_tac [cmult_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1));
-by (asm_simp_tac (ZF_ss addsimps [InfCard_csquare_eq]) 1);
+by (asm_simp_tac (!simpset addsimps [InfCard_csquare_eq]) 1);
qed "InfCard_le_cmult_eq";
(*Corollary 10.13 (1), for cardinal multiplication*)
@@ -556,14 +556,14 @@
by (resolve_tac [Un_commute RS ssubst] 1);
by (ALLGOALS
(asm_simp_tac
- (ZF_ss addsimps [InfCard_is_Limit RS Limit_has_0, InfCard_le_cmult_eq,
+ (!simpset addsimps [InfCard_is_Limit RS Limit_has_0, InfCard_le_cmult_eq,
subset_Un_iff2 RS iffD1, le_imp_subset])));
qed "InfCard_cmult_eq";
(*This proof appear to be the simplest!*)
goal CardinalArith.thy "!!K. InfCard(K) ==> K |+| K = K";
by (asm_simp_tac
- (ZF_ss addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1);
+ (!simpset addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1);
by (rtac InfCard_le_cmult_eq 1);
by (typechk_tac [Ord_0, le_refl, leI]);
by (typechk_tac [InfCard_is_Limit, Limit_has_0, Limit_has_succ]);
@@ -576,7 +576,7 @@
REPEAT (ares_tac [cadd_le_self, InfCard_is_Card] 2));
by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1);
by (resolve_tac [cadd_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1));
-by (asm_simp_tac (ZF_ss addsimps [InfCard_cdouble_eq]) 1);
+by (asm_simp_tac (!simpset addsimps [InfCard_cdouble_eq]) 1);
qed "InfCard_le_cadd_eq";
goal CardinalArith.thy
@@ -587,7 +587,7 @@
by (resolve_tac [Un_commute RS ssubst] 1);
by (ALLGOALS
(asm_simp_tac
- (ZF_ss addsimps [InfCard_le_cadd_eq,
+ (!simpset addsimps [InfCard_le_cadd_eq,
subset_Un_iff2 RS iffD1, le_imp_subset])));
qed "InfCard_cadd_eq";
@@ -601,14 +601,14 @@
goalw CardinalArith.thy [jump_cardinal_def] "Ord(jump_cardinal(K))";
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1);
-by (fast_tac (ZF_cs addSIs [Ord_ordertype]) 2);
+by (fast_tac (!claset addSIs [Ord_ordertype]) 2);
by (rewtac Transset_def);
by (safe_tac subset_cs);
-by (asm_full_simp_tac (ZF_ss addsimps [ordertype_pred_unfold]) 1);
-by (safe_tac ZF_cs);
+by (asm_full_simp_tac (!simpset addsimps [ordertype_pred_unfold]) 1);
+by (safe_tac (!claset));
by (rtac UN_I 1);
by (rtac ReplaceI 2);
-by (ALLGOALS (fast_tac (ZF_cs addSEs [well_ord_subset, predE])));
+by (ALLGOALS (fast_tac (!claset addSEs [well_ord_subset, predE])));
qed "Ord_jump_cardinal";
(*Allows selective unfolding. Less work than deriving intro/elim rules*)
@@ -624,8 +624,8 @@
by (resolve_tac [jump_cardinal_iff RS iffD2] 1);
by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel]));
by (rtac subset_refl 2);
-by (asm_simp_tac (ZF_ss addsimps [Memrel_def, subset_iff]) 1);
-by (asm_simp_tac (ZF_ss addsimps [ordertype_Memrel]) 1);
+by (asm_simp_tac (!simpset addsimps [Memrel_def, subset_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [ordertype_Memrel]) 1);
qed "K_lt_jump_cardinal";
(*The proof by contradiction: the bijection f yields a wellordering of X
@@ -643,7 +643,7 @@
by (etac (bij_is_inj RS well_ord_rvimage) 1);
by (rtac (Ord_jump_cardinal RS well_ord_Memrel) 1);
by (asm_simp_tac
- (ZF_ss addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage),
+ (!simpset addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage),
ordertype_Memrel, Ord_jump_cardinal]) 1);
qed "Card_jump_cardinal_lemma";
@@ -651,7 +651,7 @@
goal CardinalArith.thy "Card(jump_cardinal(K))";
by (rtac (Ord_jump_cardinal RS CardI) 1);
by (rewtac eqpoll_def);
-by (safe_tac (ZF_cs addSDs [ltD, jump_cardinal_iff RS iffD1]));
+by (safe_tac (!claset addSDs [ltD, jump_cardinal_iff RS iffD1]));
by (REPEAT (ares_tac [Card_jump_cardinal_lemma RS mem_irrefl] 1));
qed "Card_jump_cardinal";
@@ -696,12 +696,12 @@
goal CardinalArith.thy
"!!K' K. [| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K";
by (asm_simp_tac
- (ZF_ss addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1);
+ (!simpset addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1);
qed "Card_lt_csucc_iff";
goalw CardinalArith.thy [InfCard_def]
"!!K. InfCard(K) ==> InfCard(csucc(K))";
-by (asm_simp_tac (ZF_ss addsimps [Card_csucc, Card_is_Ord,
+by (asm_simp_tac (!simpset addsimps [Card_csucc, Card_is_Ord,
lt_csucc RS leI RSN (2,le_trans)]) 1);
qed "InfCard_csucc";
@@ -711,8 +711,8 @@
goal CardinalArith.thy
"!!n. n: nat ==> ALL A. A eqpoll n --> A : Fin(A)";
by (etac nat_induct 1);
-by (simp_tac (ZF_ss addsimps (eqpoll_0_iff::Fin.intrs)) 1);
-by (step_tac ZF_cs 1);
+by (simp_tac (!simpset addsimps (eqpoll_0_iff::Fin.intrs)) 1);
+by (step_tac (!claset) 1);
by (subgoal_tac "EX u. u:A" 1);
by (etac exE 1);
by (resolve_tac [Diff_sing_eqpoll RS revcut_rl] 1);
@@ -721,28 +721,28 @@
by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
by (assume_tac 1);
by (resolve_tac [Fin.consI] 1);
-by (fast_tac ZF_cs 1);
-by (deepen_tac (ZF_cs addIs [Fin_mono RS subsetD]) 0 1); (*SLOW*)
+by (Fast_tac 1);
+by (fast_tac (!claset addIs [subset_consI RS Fin_mono RS subsetD]) 1);
(*Now for the lemma assumed above*)
by (rewtac eqpoll_def);
-by (fast_tac (ZF_cs addSEs [bij_converse_bij RS bij_is_fun RS apply_type]) 1);
+by (best_tac (!claset addSEs [bij_converse_bij RS bij_is_fun RS apply_type]) 1);
val lemma = result();
goalw CardinalArith.thy [Finite_def] "!!A. Finite(A) ==> A : Fin(A)";
-by (fast_tac (ZF_cs addIs [lemma RS spec RS mp]) 1);
+by (fast_tac (!claset addIs [lemma RS spec RS mp]) 1);
qed "Finite_into_Fin";
goal CardinalArith.thy "!!A. A : Fin(U) ==> Finite(A)";
-by (fast_tac (ZF_cs addSIs [Finite_0, Finite_cons] addEs [Fin.induct]) 1);
+by (fast_tac (!claset addSIs [Finite_0, Finite_cons] addEs [Fin.induct]) 1);
qed "Fin_into_Finite";
goal CardinalArith.thy "Finite(A) <-> A : Fin(A)";
-by (fast_tac (ZF_cs addIs [Finite_into_Fin] addEs [Fin_into_Finite]) 1);
+by (fast_tac (!claset addIs [Finite_into_Fin] addEs [Fin_into_Finite]) 1);
qed "Finite_Fin_iff";
goal CardinalArith.thy
"!!A. [| Finite(A); Finite(B) |] ==> Finite(A Un B)";
-by (fast_tac (ZF_cs addSIs [Fin_into_Finite, Fin_UnI]
+by (fast_tac (!claset addSIs [Fin_into_Finite, Fin_UnI]
addSDs [Finite_into_Fin]
addSEs [Un_upper1 RS Fin_mono RS subsetD,
Un_upper2 RS Fin_mono RS subsetD]) 1);
@@ -754,10 +754,10 @@
goal CardinalArith.thy
"!!A. A: Fin(U) ==> x~:A --> ~ cons(x,A) lepoll A";
by (etac Fin_induct 1);
-by (simp_tac (ZF_ss addsimps [lepoll_0_iff]) 1);
+by (simp_tac (!simpset addsimps [lepoll_0_iff]) 1);
by (subgoal_tac "cons(x,cons(xa,y)) = cons(xa,cons(x,y))" 1);
-by (asm_simp_tac ZF_ss 1);
-by (fast_tac (ZF_cs addSDs [cons_lepoll_consD]) 1);
+by (Asm_simp_tac 1);
+by (fast_tac (!claset addSDs [cons_lepoll_consD]) 1);
by (fast_tac eq_cs 1);
qed "Fin_imp_not_cons_lepoll";
@@ -766,18 +766,18 @@
by (rewtac cardinal_def);
by (rtac Least_equality 1);
by (fold_tac [cardinal_def]);
-by (simp_tac (ZF_ss addsimps [succ_def]) 1);
-by (fast_tac (ZF_cs addIs [cons_eqpoll_cong, well_ord_cardinal_eqpoll]
+by (simp_tac (!simpset addsimps [succ_def]) 1);
+by (fast_tac (!claset addIs [cons_eqpoll_cong, well_ord_cardinal_eqpoll]
addSEs [mem_irrefl]
addSDs [Finite_imp_well_ord]) 1);
-by (fast_tac (ZF_cs addIs [Ord_succ, Card_cardinal, Card_is_Ord]) 1);
+by (fast_tac (!claset addIs [Ord_succ, Card_cardinal, Card_is_Ord]) 1);
by (rtac notI 1);
by (resolve_tac [Finite_into_Fin RS Fin_imp_not_cons_lepoll RS mp RS notE] 1);
by (assume_tac 1);
by (assume_tac 1);
by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans] 1);
by (eresolve_tac [le_imp_lepoll RS lepoll_trans] 1);
-by (fast_tac (ZF_cs addIs [well_ord_cardinal_eqpoll RS eqpoll_imp_lepoll]
+by (fast_tac (!claset addIs [well_ord_cardinal_eqpoll RS eqpoll_imp_lepoll]
addSDs [Finite_imp_well_ord]) 1);
qed "Finite_imp_cardinal_cons";
@@ -785,14 +785,14 @@
goal CardinalArith.thy "!!a A. [| Finite(A); a:A |] ==> succ(|A-{a}|) = |A|";
by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
by (assume_tac 1);
-by (asm_simp_tac (ZF_ss addsimps [Finite_imp_cardinal_cons,
+by (asm_simp_tac (!simpset addsimps [Finite_imp_cardinal_cons,
Diff_subset RS subset_Finite]) 1);
-by (asm_simp_tac (ZF_ss addsimps [cons_Diff]) 1);
+by (asm_simp_tac (!simpset addsimps [cons_Diff]) 1);
qed "Finite_imp_succ_cardinal_Diff";
goal CardinalArith.thy "!!a A. [| Finite(A); a:A |] ==> |A-{a}| < |A|";
by (rtac succ_leE 1);
-by (asm_simp_tac (ZF_ss addsimps [Finite_imp_succ_cardinal_Diff,
+by (asm_simp_tac (!simpset addsimps [Finite_imp_succ_cardinal_Diff,
Ord_cardinal RS le_refl]) 1);
qed "Finite_imp_cardinal_Diff";
@@ -808,11 +808,11 @@
well_ord_radd RS well_ord_cardinal_eqpoll)) RS eqpoll_sym] 1
THEN (assume_tac 1));
by (eresolve_tac [nat_cadd_eq_add RS subst] 1 THEN (assume_tac 1));
-by (asm_full_simp_tac (ZF_ss addsimps [cadd_def, eqpoll_refl]) 1);
+by (asm_full_simp_tac (!simpset addsimps [cadd_def, eqpoll_refl]) 1);
qed "nat_sum_eqpoll_sum";
goal Nat.thy "!!m. [| m le n; n:nat |] ==> m:nat";
-by (fast_tac (ZF_cs addSDs [nat_succI RS (Ord_nat RSN (2, OrdmemD))]
+by (fast_tac (!claset addSDs [nat_succI RS (Ord_nat RSN (2, OrdmemD))]
addSEs [ltE]) 1);
qed "le_in_nat";
--- a/src/ZF/Cardinal_AC.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Cardinal_AC.ML Fri Jan 03 15:01:55 1997 +0100
@@ -27,13 +27,13 @@
qed "cardinal_eqE";
goal Cardinal_AC.thy "|X| = |Y| <-> X eqpoll Y";
-by (fast_tac (ZF_cs addSEs [cardinal_cong, cardinal_eqE]) 1);
+by (fast_tac (!claset addSEs [cardinal_cong, cardinal_eqE]) 1);
qed "cardinal_eqpoll_iff";
goal Cardinal_AC.thy
"!!A. [| |A|=|B|; |C|=|D|; A Int C = 0; B Int D = 0 |] ==> \
\ |A Un C| = |B Un D|";
-by (asm_full_simp_tac (ZF_ss addsimps [cardinal_eqpoll_iff,
+by (asm_full_simp_tac (!simpset addsimps [cardinal_eqpoll_iff,
eqpoll_disjoint_Un]) 1);
qed "cardinal_disjoint_Un";
@@ -92,12 +92,12 @@
goalw Cardinal_AC.thy [surj_def] "!!f. f: surj(X,Y) ==> EX g. g: inj(Y,X)";
by (etac CollectE 1);
by (res_inst_tac [("A1", "Y"), ("B1", "%y. f-``{y}")] (AC_Pi RS exE) 1);
-by (fast_tac (ZF_cs addSEs [apply_Pair]) 1);
+by (fast_tac (!claset addSEs [apply_Pair]) 1);
by (rtac exI 1);
by (rtac f_imp_injective 1);
by (rtac Pi_type 1 THEN assume_tac 1);
-by (fast_tac (ZF_cs addDs [apply_type] addDs [Pi_memberD]) 1);
-by (fast_tac (ZF_cs addDs [apply_type] addEs [apply_equality]) 1);
+by (fast_tac (!claset addDs [apply_type] addDs [Pi_memberD]) 1);
+by (fast_tac (!claset addDs [apply_type] addEs [apply_equality]) 1);
qed "surj_implies_inj";
(*Kunen's Lemma 10.20*)
@@ -111,10 +111,10 @@
(*Kunen's Lemma 10.21*)
goal Cardinal_AC.thy
"!!K. [| InfCard(K); ALL i:K. |X(i)| le K |] ==> |UN i:K. X(i)| le K";
-by (asm_full_simp_tac (ZF_ss addsimps [InfCard_is_Card, le_Card_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [InfCard_is_Card, le_Card_iff]) 1);
by (rtac lepoll_trans 1);
by (resolve_tac [InfCard_square_eq RS eqpoll_imp_lepoll] 2);
-by (asm_simp_tac (ZF_ss addsimps [InfCard_is_Card, Card_cardinal_eq]) 2);
+by (asm_simp_tac (!simpset addsimps [InfCard_is_Card, Card_cardinal_eq]) 2);
by (rewtac lepoll_def);
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1);
by (etac (AC_ball_Pi RS exE) 1);
@@ -122,18 +122,18 @@
(*Lemma needed in both subgoals, for a fixed z*)
by (subgoal_tac
"ALL z: (UN i:K. X(i)). z: X(LEAST i. z:X(i)) & (LEAST i. z:X(i)) : K" 1);
-by (fast_tac (ZF_cs addSIs [Least_le RS lt_trans1 RS ltD, ltI]
+by (fast_tac (!claset addSIs [Least_le RS lt_trans1 RS ltD, ltI]
addSEs [LeastI, Ord_in_Ord]) 2);
by (res_inst_tac [("c", "%z. <LEAST i. z:X(i), f ` (LEAST i. z:X(i)) ` z>"),
("d", "%<i,j>. converse(f`i) ` j")]
lam_injective 1);
(*Instantiate the lemma proved above*)
by (ALLGOALS ball_tac);
-by (fast_tac (ZF_cs addEs [inj_is_fun RS apply_type]
+by (fast_tac (!claset addEs [inj_is_fun RS apply_type]
addDs [apply_type]) 1);
by (dtac apply_type 1);
by (etac conjunct2 1);
-by (asm_simp_tac (ZF_ss addsimps [left_inverse]) 1);
+by (asm_simp_tac (!simpset addsimps [left_inverse]) 1);
qed "cardinal_UN_le";
(*The same again, using csucc*)
@@ -141,7 +141,7 @@
"!!K. [| InfCard(K); ALL i:K. |X(i)| < csucc(K) |] ==> \
\ |UN i:K. X(i)| < csucc(K)";
by (asm_full_simp_tac
- (ZF_ss addsimps [Card_lt_csucc_iff, cardinal_UN_le,
+ (!simpset addsimps [Card_lt_csucc_iff, cardinal_UN_le,
InfCard_is_Card, Card_cardinal]) 1);
qed "cardinal_UN_lt_csucc";
@@ -152,8 +152,8 @@
\ (UN i:K. j(i)) < csucc(K)";
by (resolve_tac [cardinal_UN_lt_csucc RS Card_lt_imp_lt] 1);
by (assume_tac 1);
-by (fast_tac (ZF_cs addIs [Ord_cardinal_le RS lt_trans1] addEs [ltE]) 1);
-by (fast_tac (ZF_cs addSIs [Ord_UN] addEs [ltE]) 1);
+by (fast_tac (!claset addIs [Ord_cardinal_le RS lt_trans1] addEs [ltE]) 1);
+by (fast_tac (!claset addSIs [Ord_UN] addEs [ltE]) 1);
by (eresolve_tac [InfCard_is_Card RS Card_is_Ord RS Card_csucc] 1);
qed "cardinal_UN_Ord_lt_csucc";
@@ -176,7 +176,7 @@
by (res_inst_tac [("x1", "f`x")] (UN_upper RSN (2,subset_trans)) 1);
by (eresolve_tac [inj_is_fun RS apply_type] 2 THEN assume_tac 2);
by (asm_simp_tac
- (ZF_ss addsimps [inj_is_fun RS apply_rangeI, left_inverse]) 1);
+ (!simpset addsimps [inj_is_fun RS apply_rangeI, left_inverse]) 1);
val inj_UN_subset = result();
(*Simpler to require |W|=K; we'd have a bijection; but the theorem would
@@ -186,16 +186,16 @@
\ (UN w:W. j(w)) < csucc(K)";
by (excluded_middle_tac "W=0" 1);
by (asm_simp_tac (*solve the easy 0 case*)
- (ZF_ss addsimps [UN_0, InfCard_is_Card, Card_is_Ord RS Card_csucc,
+ (!simpset addsimps [UN_0, InfCard_is_Card, Card_is_Ord RS Card_csucc,
Card_is_Ord, Ord_0_lt_csucc]) 2);
by (asm_full_simp_tac
- (ZF_ss addsimps [InfCard_is_Card, le_Card_iff, lepoll_def]) 1);
+ (!simpset addsimps [InfCard_is_Card, le_Card_iff, lepoll_def]) 1);
by (safe_tac eq_cs);
by (swap_res_tac [[inj_UN_subset, cardinal_UN_Ord_lt_csucc]
MRS lt_subset_trans] 1);
by (REPEAT (assume_tac 1));
-by (fast_tac (ZF_cs addSIs [Ord_UN] addEs [ltE]) 2);
-by (asm_simp_tac (ZF_ss addsimps [inj_converse_fun RS apply_type]
+by (fast_tac (!claset addSIs [Ord_UN] addEs [ltE]) 2);
+by (asm_simp_tac (!simpset addsimps [inj_converse_fun RS apply_type]
setloop split_tac [expand_if]) 1);
qed "le_UN_Ord_lt_csucc";
--- a/src/ZF/Coind/ECR.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Coind/ECR.ML Fri Jan 03 15:01:55 1997 +0100
@@ -16,7 +16,7 @@
\ <v_clos(x, e, ve),t>:HasTyRel";
by (rtac HasTyRel.coinduct 1);
by (rtac singletonI 1);
-by (fast_tac (ZF_cs addIs Val_ValEnv.intrs) 1);
+by (fast_tac (!claset addIs Val_ValEnv.intrs) 1);
by (rtac disjI2 1);
by (etac singletonE 1);
by (REPEAT_FIRST (resolve_tac [conjI,exI]));
@@ -42,19 +42,19 @@
addEs [htr_closE])
end;
-val htr_cs = mk_htr_cs(static_cs);
+claset := mk_htr_cs (!claset);
(* Properties of the pointwise extension to environments *)
goalw ECR.thy [hastyenv_def]
"!!ve.[| ve:ValEnv; te:TyEnv; hastyenv(ve,te); <v,t>:HasTyRel |] ==> \
\ hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (stac ve_dom_owr 1);
by (assume_tac 1);
by (etac (HasTyRel.dom_subset RS subsetD RS SigmaD1 RS ValNEE) 1);
by (stac te_dom_owr 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (rtac (excluded_middle RS disjE) 1);
by (stac ve_app_owr2 1);
by (assume_tac 1);
@@ -63,21 +63,21 @@
by (assume_tac 1);
by (dtac (ve_dom_owr RS subst) 1);
by (etac (HasTyRel.dom_subset RS subsetD RS SigmaD1 RS ValNEE) 1);
-by ((fast_tac ZF_cs 1) THEN (fast_tac ZF_cs 1));
-by (asm_simp_tac (ZF_ss addsimps [ve_app_owr1,te_app_owr1]) 1);
+by ((Fast_tac 1) THEN (Fast_tac 1));
+by (asm_simp_tac (!simpset addsimps [ve_app_owr1,te_app_owr1]) 1);
qed "hastyenv_owr";
goalw ECR.thy [isofenv_def,hastyenv_def]
"!!ve.[| ve:ValEnv; te:TyEnv; isofenv(ve,te) |] ==> hastyenv(ve,te)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (dtac bspec 1);
by (assume_tac 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (dtac HasTyRel.htr_constI 1);
by (assume_tac 2);
by (etac te_appI 1);
by (etac ve_domI 1);
-by (ALLGOALS (asm_full_simp_tac ZF_ss));
+by (ALLGOALS Asm_full_simp_tac);
qed "basic_consistency_lem";
--- a/src/ZF/Coind/Language.ML Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,35 +0,0 @@
-(* Title: ZF/Coind/Language.ML
- ID: $Id$
- Author: Jacob Frost, Cambridge University Computer Laboratory
- Copyright 1995 University of Cambridge
-*)
-
-(* ############################################################ *)
-(* General lemmas *)
-(* ############################################################ *)
-
-goal ZF.thy "!!a.~a=b ==> ~a:{b}";
-by (fast_tac ZF_cs 1);
-qed "notsingletonI";
-
-goal ZF.thy "!!a.~a:{b} ==> ~a=b";
-by (fast_tac ZF_cs 1);
-qed "notsingletonD";
-
-val prems = goal ZF.thy "[| a~:{b}; a~=b ==> P |] ==> P";
-by (cut_facts_tac prems 1);
-by (resolve_tac prems 1);
-by (fast_tac ZF_cs 1);
-qed "notsingletonE";
-
-goal ZF.thy "!!x. x : A Un B ==> x: B Un A";
-by (fast_tac ZF_cs 1);
-qed "lem_fix";
-
-goal ZF.thy "!!A.[| x~:A; x=y |] ==> y~:A";
-by (fast_tac ZF_cs 1);
-qed "map_lem1";
-
-
-open Language;
-
--- a/src/ZF/Coind/MT.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Coind/MT.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,38 +11,32 @@
(* The Consistency theorem *)
(* ############################################################ *)
-val prems = goal MT.thy
- "[| c:Const; hastyenv(ve,te);<te,e_const(c),t>:ElabRel |] ==> \
-\ <v_const(c), t> : HasTyRel";
-by (cut_facts_tac prems 1);
-by (fast_tac htr_cs 1);
+goal MT.thy
+ "!!t. [| c:Const; hastyenv(ve,te);<te,e_const(c),t>:ElabRel |] ==> \
+\ <v_const(c), t> : HasTyRel";
+by (Fast_tac 1);
qed "consistency_const";
-val prems = goalw MT.thy [hastyenv_def]
- "[| x:ve_dom(ve); hastyenv(ve,te); <te,e_var(x),t>:ElabRel |] ==> \
+goalw MT.thy [hastyenv_def]
+ "!!t. [| x:ve_dom(ve); hastyenv(ve,te); <te,e_var(x),t>:ElabRel |] ==> \
\ <ve_app(ve,x),t>:HasTyRel";
-by (cut_facts_tac prems 1);
-by (fast_tac static_cs 1);
+by (Fast_tac 1);
qed "consistency_var";
-val prems = goalw MT.thy [hastyenv_def]
- "[| ve:ValEnv; x:ExVar; e:Exp; hastyenv(ve,te); \
-\ <te,e_fn(x,e),t>:ElabRel \
-\ |] ==> \
-\ <v_clos(x, e, ve), t> : HasTyRel";
-by (cut_facts_tac prems 1);
-by (best_tac htr_cs 1);
+goalw MT.thy [hastyenv_def]
+ "!!t. [| ve:ValEnv; x:ExVar; e:Exp; hastyenv(ve,te); \
+\ <te,e_fn(x,e),t>:ElabRel \
+\ |] ==> <v_clos(x, e, ve), t> : HasTyRel";
+by (Best_tac 1);
qed "consistency_fn";
-val MT_cs = ZF_cs
- addIs (Ty.intrs@TyEnv.intrs@Val_ValEnv.intrs)
- addDs [te_owrE,(ElabRel.dom_subset RS subsetD)];
+AddIs (Ty.intrs@TyEnv.intrs@Val_ValEnv.intrs);
+AddDs [te_owrE,(ElabRel.dom_subset RS subsetD)];
-val MT_ss =
- ZF_ss addsimps
- [ve_dom_owr,te_dom_owr,ve_app_owr1,ve_app_owr2,te_app_owr1,te_app_owr2];
+Addsimps [ve_dom_owr, te_dom_owr, ve_app_owr1, ve_app_owr2,
+ te_app_owr1, te_app_owr2];
val clean_tac =
REPEAT_FIRST (fn i =>
@@ -57,7 +51,7 @@
\ <cl,t>:HasTyRel";
by (cut_facts_tac prems 1);
by (etac elab_fixE 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (EVERY [forward_tac [subst] 1,atac 2,rtac htr_closCI 1]);
by clean_tac;
by (rtac ve_owrI 1);
@@ -70,12 +64,12 @@
by (assume_tac 1);
by (rtac ElabRel.elab_fnI 1);
by clean_tac;
-by (asm_simp_tac MT_ss 1);
+by (Asm_simp_tac 1);
by (stac ve_dom_owr 1);
by (assume_tac 1);
by (etac subst 1);
by (rtac v_closNE 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (rtac PowI 1);
by (stac ve_dom_owr 1);
@@ -84,15 +78,8 @@
by (rtac v_closNE 1);
by (rtac subsetI 1);
by (etac RepFunE 1);
-by (dtac lem_fix 1);
-by (etac UnE' 1);
-by (rtac UnI1 1);
-by (etac singletonE 1);
-by (asm_full_simp_tac MT_ss 1);
-by (rtac UnI2 1);
-by (etac notsingletonE 1);
-by (dtac not_sym 1);
-by (asm_full_simp_tac MT_ss 1);
+by (excluded_middle_tac "f=y" 1);
+by (Auto_tac());
qed "consistency_fix";
@@ -151,7 +138,7 @@
by (assume_tac 1);
by (assume_tac 2);
by (rewrite_tac [hastyenv_def]);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "consistency_app2";
fun mt_cases_tac t = ((rtac t 1) THEN (TRYALL assume_tac));
@@ -180,7 +167,7 @@
by (cut_facts_tac prems 1);
by (rtac (htr_constE) 1);
by (dtac consistency 1);
-by (fast_tac (ZF_cs addSIs [basic_consistency_lem]) 1);
+by (fast_tac (!claset addSIs [basic_consistency_lem]) 1);
by (assume_tac 1);
qed "basic_consistency";
--- a/src/ZF/Coind/Map.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Coind/Map.ML Fri Jan 03 15:01:55 1997 +0100
@@ -38,7 +38,7 @@
by (rtac ([Sigma_mono, product_univ] MRS subset_trans) 1);
by (etac subset_trans 1);
by (rtac (arg_subset_eclose RS univ_mono) 1);
-by (simp_tac (ZF_ss addsimps [Union_Pow_eq]) 1);
+by (simp_tac (!simpset addsimps [Union_Pow_eq]) 1);
qed "MapQU_lemma";
(* Theorems *)
@@ -49,7 +49,7 @@
by (rtac (MapQU_lemma RS subsetD) 1);
by (rtac subsetI 1);
by (eresolve_tac prems 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "mapQU";
(* ############################################################ *)
@@ -57,7 +57,7 @@
(* ############################################################ *)
goalw Map.thy [PMap_def,TMap_def] "!!A.B<=C ==> PMap(A,B)<=PMap(A,C)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "map_mono";
(* Rename to pmap_mono *)
@@ -69,7 +69,7 @@
(** map_emp **)
goalw Map.thy [map_emp_def,PMap_def,TMap_def] "map_emp:PMap(A,B)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac image_02 1);
qed "pmap_empI";
@@ -78,21 +78,21 @@
goalw Map.thy [map_owr_def,PMap_def,TMap_def]
"!! A.[| m:PMap(A,B); a:A; b:B |] ==> map_owr(m,a,b):PMap(A,B)";
-by (safe_tac ZF_cs);
-by (ALLGOALS (asm_full_simp_tac (ZF_ss addsimps [if_iff])));
-by (fast_tac ZF_cs 1);
-by (fast_tac ZF_cs 1);
-by (fast_tac ZF_cs 1);
-
+by (Step_tac 1);
+by (ALLGOALS (asm_full_simp_tac (!simpset addsimps [if_iff])));
+by (Fast_tac 1);
+by (Fast_tac 1);
+by (Deepen_tac 2 1);
+(*Remaining subgoal*)
by (rtac (excluded_middle RS disjE) 1);
by (etac image_Sigma1 1);
by (dres_inst_tac [("psi", "?uu ~: B")] asm_rl 1);
by (asm_full_simp_tac
- (ZF_ss addsimps [qbeta] setloop split_tac [expand_if]) 1);
-by (safe_tac FOL_cs);
+ (!simpset addsimps [qbeta] setloop split_tac [expand_if]) 1);
+by (safe_tac (!claset));
by (dres_inst_tac [("psi", "?uu ~: B")] asm_rl 3);
-by (ALLGOALS (asm_full_simp_tac ZF_ss));
-by (fast_tac ZF_cs 1);
+by (ALLGOALS Asm_full_simp_tac);
+by (Fast_tac 1);
qed "pmap_owrI";
(** map_app **)
@@ -101,7 +101,7 @@
"!!m.[| m:TMap(A,B); a:domain(m) |] ==> map_app(m,a) ~=0";
by (etac domainE 1);
by (dtac imageI 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (etac not_emptyI 1);
qed "tmap_app_notempty";
@@ -116,7 +116,7 @@
by (assume_tac 1);
by (dtac tmap_appI 1);
by (assume_tac 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "pmap_appI";
(** domain **)
@@ -144,7 +144,7 @@
qed "domain_UN";
goal Map.thy "domain(Sigma(A,B)) = {x:A.EX y.y:B(x)}";
-by (simp_tac (ZF_ss addsimps [domain_UN,domain_0,domain_cons]) 1);
+by (simp_tac (!simpset addsimps [domain_UN,domain_0,domain_cons]) 1);
by (fast_tac eq_cs 1);
qed "domain_Sigma";
@@ -156,7 +156,7 @@
goalw Map.thy [map_owr_def]
"!!a.b ~= 0 ==> domain(map_owr(f,a,b)) = {a} Un domain(f)";
-by (simp_tac (if_ss addsimps [domain_Sigma]) 1);
+by (simp_tac (!simpset addsimps [domain_Sigma]) 1);
by (rtac equalityI 1);
by (fast_tac eq_cs 1);
by (rtac subsetI 1);
@@ -164,11 +164,9 @@
by (assume_tac 1);
by (etac UnE' 1);
by (etac singletonE 1);
-by (asm_simp_tac if_ss 1);
+by (Asm_simp_tac 1);
by (fast_tac eq_cs 1);
-by (etac notsingletonE 1);
-by (asm_simp_tac if_ss 1);
-by (fast_tac eq_cs 1);
+by (fast_tac (eq_cs addss (!simpset)) 1);
qed "map_domain_owr";
(** Application **)
@@ -176,8 +174,8 @@
goalw Map.thy [map_app_def,map_owr_def]
"map_app(map_owr(f,a,b),a) = b";
by (stac qbeta 1);
-by (fast_tac ZF_cs 1);
-by (simp_tac if_ss 1);
+by (Fast_tac 1);
+by (Simp_tac 1);
qed "map_app_owr1";
goalw Map.thy [map_app_def,map_owr_def]
@@ -187,7 +185,7 @@
by (assume_tac 1);
by (fast_tac eq_cs 1);
by (etac (qbeta RS ssubst) 1);
-by (asm_simp_tac if_ss 1);
+by (Asm_simp_tac 1);
qed "map_app_owr2";
--- a/src/ZF/Coind/Static.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Coind/Static.ML Fri Jan 03 15:01:55 1997 +0100
@@ -21,20 +21,10 @@
val elab_appE =
ElabRel.mk_cases Exp.con_defs "<te,e_app(e1,e2),t>:ElabRel";
-fun mk_static_cs cs=
- let open ElabRel in
- ( cs
- addSIs [elab_constI,elab_varI,elab_fnI,elab_fixI]
- addSEs [elab_constE,elab_varE,elab_fixE]
- addIs [elab_appI]
- addEs [elab_appE,elab_fnE]
- addDs [ElabRel.dom_subset RS subsetD]
- ) end ;
-
-val static_cs = mk_static_cs ZF_cs;
-
-
-
-
-
-
+let open ElabRel in
+claset := !claset addSIs [elab_constI,elab_varI,elab_fnI,elab_fixI]
+ addSEs [elab_constE,elab_varE,elab_fixE]
+ addIs [elab_appI]
+ addEs [elab_appE,elab_fnE]
+ addDs [ElabRel.dom_subset RS subsetD]
+end;
--- a/src/ZF/Coind/Types.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Coind/Types.ML Fri Jan 03 15:01:55 1997 +0100
@@ -15,7 +15,7 @@
goal Types.thy "te_rec(te_emp,c_te_emp,f_te_owr) = c_te_emp";
by (rtac (te_rec_def RS def_Vrec RS trans) 1);
-by (simp_tac (ZF_ss addsimps TyEnv.case_eqns) 1);
+by (simp_tac (!simpset addsimps (rank_te_owr1::TyEnv.case_eqns)) 1);
qed "te_rec_emp";
goal Types.thy
@@ -26,38 +26,36 @@
qed "te_rec_owr";
goalw Types.thy [te_dom_def] "te_dom(te_emp) = 0";
-by (simp_tac (ZF_ss addsimps [te_rec_emp]) 1);
+by (simp_tac (!simpset addsimps [te_rec_emp]) 1);
qed "te_dom_emp";
goalw Types.thy [te_dom_def] "te_dom(te_owr(te,x,v)) = te_dom(te) Un {x}";
-by (simp_tac (ZF_ss addsimps [te_rec_owr]) 1);
+by (simp_tac (!simpset addsimps [te_rec_owr]) 1);
qed "te_dom_owr";
goalw Types.thy [te_app_def] "te_app(te_owr(te,x,t),x) = t";
-by (simp_tac (ZF_ss addsimps [te_rec_owr]) 1);
+by (simp_tac (!simpset addsimps [te_rec_owr]) 1);
qed "te_app_owr1";
-val prems = goalw Types.thy [te_app_def]
- "x ~= y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)";
-by (cut_facts_tac prems 1);
-by (asm_simp_tac (ZF_ss addsimps [te_rec_owr,(not_sym RS if_not_P)]) 1);
+goalw Types.thy [te_app_def]
+ "!!x y. x ~= y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)";
+by (asm_simp_tac (!simpset addsimps [te_rec_owr,(not_sym RS if_not_P)]) 1);
qed "te_app_owr2";
-val prems = goal Types.thy
- "[| te:TyEnv; x:ExVar; x:te_dom(te) |] ==> te_app(te,x):Ty";
-by (cut_facts_tac prems 1);
+goal Types.thy
+ "!!te. [| te:TyEnv; x:ExVar; x:te_dom(te) |] ==> te_app(te,x):Ty";
by (res_inst_tac [("P","x:te_dom(te)")] impE 1);
by (assume_tac 2);
by (assume_tac 2);
by (etac TyEnv.induct 1);
-by (simp_tac (ZF_ss addsimps [te_dom_emp]) 1);
+by (simp_tac (!simpset addsimps [te_dom_emp]) 1);
by (rtac impI 1);
by (rtac (excluded_middle RS disjE) 1);
by (stac te_app_owr2 1);
by (assume_tac 1);
-by (asm_full_simp_tac (ZF_ss addsimps [te_dom_owr]) 1);
-by (fast_tac ZF_cs 1);
-by (asm_simp_tac (ZF_ss addsimps [te_app_owr1]) 1);
+by (asm_full_simp_tac (!simpset addsimps [te_dom_owr]) 1);
+by (Fast_tac 1);
+by (asm_simp_tac (!simpset addsimps [te_app_owr1]) 1);
qed "te_appI";
--- a/src/ZF/Coind/Values.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Coind/Values.ML Fri Jan 03 15:01:55 1997 +0100
@@ -16,7 +16,7 @@
by (etac Val_ValEnv.elim 1);
by (eresolve_tac prems 3);
by (rewrite_tac Val_ValEnv.con_defs);
-by (ALLGOALS (fast_tac qsum_cs));
+by (ALLGOALS Fast_tac);
qed "ValEnvE";
val prems = goalw Values.thy (Part_def::[hd (tl Val_ValEnv.defs)])
@@ -37,7 +37,7 @@
by (assume_tac 1);
by (assume_tac 1);
by (rewrite_tac Val_ValEnv.con_defs);
-by (fast_tac qsum_cs 1);
+by (Fast_tac 1);
qed "ValE";
(* Nonempty sets *)
@@ -45,8 +45,8 @@
val prems = goal Values.thy "A ~= 0 ==> EX a.a:A";
by (cut_facts_tac prems 1);
by (rtac (foundation RS disjE) 1);
-by (fast_tac ZF_cs 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
+by (Fast_tac 1);
qed "set_notemptyE";
goalw Values.thy (sum_def::QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs)
@@ -56,7 +56,7 @@
by (rtac SigmaI 1);
by (rtac singletonI 1);
by (rtac UnI1 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "v_closNE";
goalw Values.thy (sum_def::QPair_def::QInl_def::QInr_def::Val_ValEnv.con_defs)
@@ -88,7 +88,7 @@
"[| ve:ValEnv; v ~=0 |] ==> ve_dom(ve_owr(ve,x,v)) = ve_dom(ve) Un {x}";
by (cut_facts_tac prems 1);
by (etac ValEnvE 1);
-by (asm_full_simp_tac (ZF_ss addsimps Val_ValEnv.case_eqns) 1);
+by (asm_full_simp_tac (!simpset addsimps Val_ValEnv.case_eqns) 1);
by (stac map_domain_owr 1);
by (assume_tac 1);
by (rtac Un_commute 1);
@@ -97,16 +97,16 @@
goalw Values.thy [ve_app_def,ve_owr_def]
"!!ve. ve:ValEnv ==> ve_app(ve_owr(ve,x,v),x) = v";
by (etac ValEnvE 1);
-by (asm_full_simp_tac (ZF_ss addsimps Val_ValEnv.case_eqns) 1);
+by (asm_full_simp_tac (!simpset addsimps Val_ValEnv.case_eqns) 1);
by (rtac map_app_owr1 1);
qed "ve_app_owr1";
goalw Values.thy [ve_app_def,ve_owr_def]
"!!ve.ve:ValEnv ==> x ~= y ==> ve_app(ve_owr(ve,x,v),y) = ve_app(ve,y)";
by (etac ValEnvE 1);
-by (asm_full_simp_tac (ZF_ss addsimps Val_ValEnv.case_eqns) 1);
+by (asm_full_simp_tac (!simpset addsimps Val_ValEnv.case_eqns) 1);
by (rtac map_app_owr2 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "ve_app_owr2";
(* Introduction rules for operators on value environments *)
@@ -115,7 +115,7 @@
"!!ve.[| ve:ValEnv; x:ve_dom(ve) |] ==> ve_app(ve,x):Val";
by (etac ValEnvE 1);
by (hyp_subst_tac 1);
-by (asm_full_simp_tac (ZF_ss addsimps Val_ValEnv.case_eqns) 1);
+by (asm_full_simp_tac (!simpset addsimps Val_ValEnv.case_eqns) 1);
by (rtac pmap_appI 1);
by (assume_tac 1);
by (assume_tac 1);
@@ -125,7 +125,7 @@
"!!ve.[| ve:ValEnv; x:ve_dom(ve) |] ==> x:ExVar";
by (etac ValEnvE 1);
by (hyp_subst_tac 1);
-by (asm_full_simp_tac (ZF_ss addsimps Val_ValEnv.case_eqns) 1);
+by (asm_full_simp_tac (!simpset addsimps Val_ValEnv.case_eqns) 1);
by (rtac pmap_domainD 1);
by (assume_tac 1);
by (assume_tac 1);
@@ -140,7 +140,7 @@
"!!ve.[|ve:ValEnv; x:ExVar; v:Val |] ==> ve_owr(ve,x,v):ValEnv";
by (etac ValEnvE 1);
by (hyp_subst_tac 1);
-by (asm_full_simp_tac (ZF_ss addsimps Val_ValEnv.case_eqns) 1);
+by (asm_full_simp_tac (!simpset addsimps Val_ValEnv.case_eqns) 1);
by (rtac Val_ValEnv.ve_mk_I 1);
by (etac pmap_owrI 1);
by (assume_tac 1);
--- a/src/ZF/Epsilon.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Epsilon.ML Fri Jan 03 15:01:55 1997 +0100
@@ -64,9 +64,9 @@
"!!X A n. [| Transset(X); A<=X; n: nat |] ==> \
\ nat_rec(n, A, %m r. Union(r)) <= X";
by (etac nat_induct 1);
-by (asm_simp_tac (ZF_ss addsimps [nat_rec_0]) 1);
-by (asm_simp_tac (ZF_ss addsimps [nat_rec_succ]) 1);
-by (fast_tac ZF_cs 1);
+by (asm_simp_tac (!simpset addsimps [nat_rec_0]) 1);
+by (asm_simp_tac (!simpset addsimps [nat_rec_succ]) 1);
+by (Fast_tac 1);
qed "eclose_least_lemma";
goalw Epsilon.thy [eclose_def]
@@ -86,7 +86,7 @@
by (etac (arg_subset_eclose RS subsetD) 2);
by (etac base 2);
by (rewtac Transset_def);
-by (fast_tac (ZF_cs addEs [step,ecloseD]) 1);
+by (fast_tac (!claset addEs [step,ecloseD]) 1);
qed "eclose_induct_down";
goal Epsilon.thy "!!X. Transset(X) ==> eclose(X) = X";
@@ -118,7 +118,7 @@
(* j : eclose(A) ==> Memrel(eclose(A)) -`` j = j *)
val under_Memrel_eclose = Transset_eclose RS under_Memrel;
-bind_thm ("wfrec_ssubst", wf_Memrel RS wfrec RS ssubst);
+val wfrec_ssubst = standard (wf_Memrel RS wfrec RS ssubst);
val [kmemj,jmemi] = goal Epsilon.thy
"[| k:eclose({j}); j:eclose({i}) |] ==> \
@@ -126,7 +126,7 @@
by (rtac (kmemj RS eclose_induct) 1);
by (rtac wfrec_ssubst 1);
by (rtac wfrec_ssubst 1);
-by (asm_simp_tac (ZF_ss addsimps [under_Memrel_eclose,
+by (asm_simp_tac (!simpset addsimps [under_Memrel_eclose,
jmemi RSN (2,mem_eclose_sing_trans)]) 1);
qed "wfrec_eclose_eq";
@@ -139,7 +139,7 @@
goalw Epsilon.thy [transrec_def]
"transrec(a,H) = H(a, lam x:a. transrec(x,H))";
by (rtac wfrec_ssubst 1);
-by (simp_tac (ZF_ss addsimps [wfrec_eclose_eq2, arg_in_eclose_sing,
+by (simp_tac (!simpset addsimps [wfrec_eclose_eq2, arg_in_eclose_sing,
under_Memrel_eclose]) 1);
qed "transrec";
@@ -178,7 +178,7 @@
(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*)
goal Epsilon.thy "rank(a) = (UN y:a. succ(rank(y)))";
by (stac (rank_def RS def_transrec) 1);
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
qed "rank";
goal Epsilon.thy "Ord(rank(a))";
@@ -192,7 +192,7 @@
val [major] = goal Epsilon.thy "Ord(i) ==> rank(i) = i";
by (rtac (major RS trans_induct) 1);
by (stac rank 1);
-by (asm_simp_tac (ZF_ss addsimps [Ord_equality]) 1);
+by (asm_simp_tac (!simpset addsimps [Ord_equality]) 1);
qed "rank_of_Ord";
goal Epsilon.thy "!!a b. a:b ==> rank(a) < rank(b)";
@@ -228,14 +228,14 @@
goal Epsilon.thy "rank(0) = 0";
by (rtac (rank RS trans) 1);
-by (fast_tac (ZF_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
qed "rank_0";
goal Epsilon.thy "rank(succ(x)) = succ(rank(x))";
by (rtac (rank RS trans) 1);
by (rtac ([UN_least, succI1 RS UN_upper] MRS equalityI) 1);
by (etac succE 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (etac (rank_lt RS leI RS succ_leI RS le_imp_subset) 1);
qed "rank_succ";
@@ -297,3 +297,32 @@
by (rtac ([Transset_eclose, subset_refl] MRS eclose_least) 1);
by (rtac arg_subset_eclose 1);
qed "eclose_idem";
+
+(*Transfinite recursion for definitions based on the three cases of ordinals.
+*)
+
+goal thy "transrec2(0,a,b) = a";
+by (rtac (transrec2_def RS def_transrec RS trans) 1);
+by (Simp_tac 1);
+qed "transrec2_0";
+
+goal thy "(THE j. i=j) = i";
+by (fast_tac (!claset addSIs [the_equality]) 1);
+qed "THE_eq";
+
+goal thy "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))";
+by (rtac (transrec2_def RS def_transrec RS trans) 1);
+by (simp_tac (!simpset addsimps [succ_not_0, THE_eq, if_P]
+ setsolver K Fast_tac) 1);
+qed "transrec2_succ";
+
+goal thy "!!i. Limit(i) ==> transrec2(i,a,b) = (UN j<i. transrec2(j,a,b))";
+by (rtac (transrec2_def RS def_transrec RS trans) 1);
+by (resolve_tac [if_not_P RS trans] 1 THEN
+ fast_tac (!claset addSDs [Limit_has_0] addSEs [ltE]) 1);
+by (resolve_tac [if_not_P RS trans] 1 THEN
+ fast_tac (!claset addSEs [succ_LimitE]) 1);
+by (Simp_tac 1);
+qed "transrec2_Limit";
+
+Addsimps [transrec2_0, transrec2_succ];
--- a/src/ZF/Epsilon.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Epsilon.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,13 +6,23 @@
Epsilon induction and recursion
*)
-Epsilon = Nat + "mono" +
-consts
- eclose,rank :: i=>i
- transrec :: [i, [i,i]=>i] =>i
+Epsilon = Nat + mono +
+constdefs
+ eclose :: i=>i
+ "eclose(A) == UN n:nat. nat_rec(n, A, %m r. Union(r))"
-defs
- eclose_def "eclose(A) == UN n:nat. nat_rec(n, A, %m r. Union(r))"
- transrec_def "transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)"
- rank_def "rank(a) == transrec(a, %x f. UN y:x. succ(f`y))"
+ transrec :: [i, [i,i]=>i] =>i
+ "transrec(a,H) == wfrec(Memrel(eclose({a})), a, H)"
+
+ rank :: i=>i
+ "rank(a) == transrec(a, %x f. UN y:x. succ(f`y))"
+
+ transrec2 :: [i, i, [i,i]=>i] =>i
+ "transrec2(k, a, b) ==
+ transrec(k,
+ %i r. if(i=0, a,
+ if(EX j. i=succ(j),
+ b(THE j. i=succ(j), r`(THE j. i=succ(j))),
+ UN j<i. r`j)))"
+
end
--- a/src/ZF/EquivClass.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/EquivClass.ML Fri Jan 03 15:01:55 1997 +0100
@@ -16,12 +16,12 @@
goalw EquivClass.thy [trans_def,sym_def]
"!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r";
-by (fast_tac (ZF_cs addSEs [converseD,compE]) 1);
+by (fast_tac (!claset) 1);
qed "sym_trans_comp_subset";
goalw EquivClass.thy [refl_def]
"!!A r. [| refl(A,r); r <= A*A |] ==> r <= converse(r) O r";
-by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 1);
+by (fast_tac (!claset) 1);
qed "refl_comp_subset";
goalw EquivClass.thy [equiv_def]
@@ -35,10 +35,7 @@
"!!A r. [| converse(r) O r = r; domain(r) = A |] ==> equiv(A,r)";
by (etac equalityE 1);
by (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r" 1);
-by (safe_tac ZF_cs);
-by (fast_tac (ZF_cs addSIs [converseI] addIs [compI]) 3);
-by (ALLGOALS (fast_tac
- (ZF_cs addSIs [converseI] addIs [compI] addSEs [compE])));
+by (ALLGOALS (fast_tac (!claset)));
qed "comp_equivI";
(** Equivalence classes **)
@@ -46,25 +43,25 @@
(*Lemma for the next result*)
goalw EquivClass.thy [trans_def,sym_def]
"!!A r. [| sym(r); trans(r); <a,b>: r |] ==> r``{a} <= r``{b}";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "equiv_class_subset";
goalw EquivClass.thy [equiv_def]
"!!A r. [| equiv(A,r); <a,b>: r |] ==> r``{a} = r``{b}";
by (safe_tac (subset_cs addSIs [equalityI, equiv_class_subset]));
by (rewtac sym_def);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "equiv_class_eq";
goalw EquivClass.thy [equiv_def,refl_def]
"!!A r. [| equiv(A,r); a: A |] ==> a: r``{a}";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "equiv_class_self";
(*Lemma for the next result*)
goalw EquivClass.thy [equiv_def,refl_def]
"!!A r. [| equiv(A,r); r``{b} <= r``{a}; b: A |] ==> <a,b>: r";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "subset_equiv_class";
val prems = goal EquivClass.thy
@@ -75,7 +72,7 @@
(*thus r``{a} = r``{b} as well*)
goalw EquivClass.thy [equiv_def,trans_def,sym_def]
"!!A r. [| equiv(A,r); x: (r``{a} Int r``{b}) |] ==> <a,b>: r";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "equiv_class_nondisjoint";
goalw EquivClass.thy [equiv_def] "!!A r. equiv(A,r) ==> r <= A*A";
@@ -84,14 +81,14 @@
goal EquivClass.thy
"!!A r. equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A";
-by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq]
- addDs [equiv_type]) 1);
+by (fast_tac (!claset addIs [eq_equiv_class, equiv_class_eq]
+ addDs [equiv_type]) 1);
qed "equiv_class_eq_iff";
goal EquivClass.thy
"!!A r. [| equiv(A,r); x: A; y: A |] ==> r``{x} = r``{y} <-> <x,y>: r";
-by (fast_tac (ZF_cs addIs [eq_equiv_class, equiv_class_eq]
- addDs [equiv_type]) 1);
+by (fast_tac (!claset addIs [eq_equiv_class, equiv_class_eq]
+ addDs [equiv_type]) 1);
qed "eq_equiv_class_iff";
(*** Quotients ***)
@@ -118,10 +115,10 @@
goalw EquivClass.thy [quotient_def]
"!!A r. [| equiv(A,r); X: A/r; Y: A/r |] ==> X=Y | (X Int Y <= 0)";
-by (safe_tac (ZF_cs addSIs [equiv_class_eq]));
+by (safe_tac (!claset addSIs [equiv_class_eq]));
by (assume_tac 1);
by (rewrite_goals_tac [equiv_def,trans_def,sym_def]);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "quotient_disj";
(**** Defining unary operations upon equivalence classes ****)
@@ -138,7 +135,7 @@
by (etac equiv_class_self 2);
by (assume_tac 2);
by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "UN_equiv_class";
(*type checking of UN x:r``{a}. b(x) *)
@@ -147,8 +144,8 @@
\ !!x. x : A ==> b(x) : B |] \
\ ==> (UN x:X. b(x)) : B";
by (cut_facts_tac prems 1);
-by (safe_tac ZF_cs);
-by (asm_simp_tac (ZF_ss addsimps (UN_equiv_class::prems)) 1);
+by (safe_tac (!claset));
+by (asm_simp_tac (!simpset addsimps (UN_equiv_class::prems)) 1);
qed "UN_equiv_class_type";
(*Sufficient conditions for injectiveness. Could weaken premises!
@@ -160,7 +157,7 @@
\ !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |] \
\ ==> X=Y";
by (cut_facts_tac prems 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac equiv_class_eq 1);
by (REPEAT (ares_tac prems 1));
by (etac box_equals 1);
@@ -173,29 +170,29 @@
goalw EquivClass.thy [congruent_def,congruent2_def,equiv_def,refl_def]
"!!A r. [| equiv(A,r); congruent2(r,b); a: A |] ==> congruent(r,b(a))";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "congruent2_implies_congruent";
val equivA::prems = goalw EquivClass.thy [congruent_def]
"[| equiv(A,r); congruent2(r,b); a: A |] ==> \
\ congruent(r, %x1. UN x2:r``{a}. b(x1,x2))";
by (cut_facts_tac (equivA::prems) 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1);
by (assume_tac 1);
-by (asm_simp_tac (ZF_ss addsimps [equivA RS UN_equiv_class,
- congruent2_implies_congruent]) 1);
+by (asm_simp_tac (!simpset addsimps [equivA RS UN_equiv_class,
+ congruent2_implies_congruent]) 1);
by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "congruent2_implies_congruent_UN";
val prems as equivA::_ = goal EquivClass.thy
"[| equiv(A,r); congruent2(r,b); a1: A; a2: A |] \
\ ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)";
by (cut_facts_tac prems 1);
-by (asm_simp_tac (ZF_ss addsimps [equivA RS UN_equiv_class,
- congruent2_implies_congruent,
- congruent2_implies_congruent_UN]) 1);
+by (asm_simp_tac (!simpset addsimps [equivA RS UN_equiv_class,
+ congruent2_implies_congruent,
+ congruent2_implies_congruent_UN]) 1);
qed "UN_equiv_class2";
(*type checking*)
@@ -205,7 +202,7 @@
\ !!x1 x2. [| x1: A; x2: A |] ==> b(x1,x2) : B \
\ |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B";
by (cut_facts_tac prems 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
congruent2_implies_congruent_UN,
congruent2_implies_congruent, quotientI]) 1));
@@ -220,7 +217,7 @@
\ !! y z w. [| w: A; <y,z> : r |] ==> b(w,y) = b(w,z) \
\ |] ==> congruent2(r,b)";
by (cut_facts_tac prems 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac trans 1);
by (REPEAT (ares_tac prems 1
ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1));
@@ -248,10 +245,10 @@
val congt' = rewrite_rule [congruent_def] congt;
by (cut_facts_tac [ZinA] 1);
by (rewtac congruent_def);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1);
by (assume_tac 1);
-by (asm_simp_tac (ZF_ss addsimps [commute,
- [equivA, congt] MRS UN_equiv_class]) 1);
+by (asm_simp_tac (!simpset addsimps [commute,
+ [equivA, congt] MRS UN_equiv_class]) 1);
by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1));
qed "congruent_commuteI";
--- a/src/ZF/Finite.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Finite.ML Fri Jan 03 15:01:55 1997 +0100
@@ -38,29 +38,31 @@
qed "Fin_induct";
(** Simplification for Fin **)
-val Fin_ss = arith_ss addsimps Fin.intrs;
+Addsimps Fin.intrs;
(*The union of two finite sets is finite.*)
-val major::prems = goal Finite.thy
- "[| b: Fin(A); c: Fin(A) |] ==> b Un c : Fin(A)";
-by (rtac (major RS Fin_induct) 1);
-by (ALLGOALS (asm_simp_tac (Fin_ss addsimps (prems@[Un_0, Un_cons]))));
+goal Finite.thy
+ "!!b c. [| b: Fin(A); c: Fin(A) |] ==> b Un c : Fin(A)";
+by (etac Fin_induct 1);
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [Un_cons])));
qed "Fin_UnI";
+Addsimps [Fin_UnI];
+
(*The union of a set of finite sets is finite.*)
val [major] = goal Finite.thy "C : Fin(Fin(A)) ==> Union(C) : Fin(A)";
by (rtac (major RS Fin_induct) 1);
-by (ALLGOALS (asm_simp_tac (Fin_ss addsimps [Union_0, Union_cons, Fin_UnI])));
+by (ALLGOALS Asm_simp_tac);
qed "Fin_UnionI";
(*Every subset of a finite set is finite.*)
goal Finite.thy "!!b A. b: Fin(A) ==> ALL z. z<=b --> z: Fin(A)";
by (etac Fin_induct 1);
-by (simp_tac (Fin_ss addsimps [subset_empty_iff]) 1);
-by (asm_simp_tac (ZF_ss addsimps subset_cons_iff::distrib_simps) 1);
-by (safe_tac ZF_cs);
+by (simp_tac (!simpset addsimps [subset_empty_iff]) 1);
+by (asm_simp_tac (!simpset addsimps subset_cons_iff::distrib_simps) 1);
+by (safe_tac (!claset));
by (eres_inst_tac [("b","z")] (cons_Diff RS subst) 1);
-by (asm_simp_tac Fin_ss 1);
+by (Asm_simp_tac 1);
qed "Fin_subset_lemma";
goal Finite.thy "!!c b A. [| c<=b; b: Fin(A) |] ==> c: Fin(A)";
@@ -74,7 +76,7 @@
\ |] ==> c<=b --> P(b-c)";
by (rtac (major RS Fin_induct) 1);
by (stac Diff_cons 2);
-by (ALLGOALS (asm_simp_tac (Fin_ss addsimps (prems@[Diff_0, cons_subset_iff,
+by (ALLGOALS (asm_simp_tac (!simpset addsimps (prems@[cons_subset_iff,
Diff_subset RS Fin_subset]))));
qed "Fin_0_induct_lemma";
@@ -91,9 +93,9 @@
(*Functions from a finite ordinal*)
val prems = goal Finite.thy "n: nat ==> n->A <= Fin(nat*A)";
by (nat_ind_tac "n" prems 1);
-by (simp_tac (ZF_ss addsimps [Pi_empty1, Fin.emptyI, subset_iff, cons_iff]) 1);
-by (asm_simp_tac (ZF_ss addsimps [succ_def, mem_not_refl RS cons_fun_eq]) 1);
-by (fast_tac (ZF_cs addSIs [Fin.consI]) 1);
+by (simp_tac (!simpset addsimps [Pi_empty1, subset_iff, cons_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [succ_def, mem_not_refl RS cons_fun_eq]) 1);
+by (fast_tac (!claset addSIs [Fin.consI]) 1);
qed "nat_fun_subset_Fin";
@@ -112,14 +114,14 @@
goal Finite.thy "!!h. h: A -||>B ==> h: domain(h) -> B";
by (etac FiniteFun.induct 1);
-by (simp_tac (ZF_ss addsimps [empty_fun, domain_0]) 1);
-by (asm_simp_tac (ZF_ss addsimps [fun_extend3, domain_cons]) 1);
+by (simp_tac (!simpset addsimps [empty_fun, domain_0]) 1);
+by (asm_simp_tac (!simpset addsimps [fun_extend3, domain_cons]) 1);
qed "FiniteFun_is_fun";
goal Finite.thy "!!h. h: A -||>B ==> domain(h) : Fin(A)";
by (etac FiniteFun.induct 1);
-by (simp_tac (Fin_ss addsimps [domain_0]) 1);
-by (asm_simp_tac (Fin_ss addsimps [domain_cons]) 1);
+by (simp_tac (!simpset addsimps [domain_0]) 1);
+by (asm_simp_tac (!simpset addsimps [domain_cons]) 1);
qed "FiniteFun_domain_Fin";
bind_thm ("FiniteFun_apply_type", FiniteFun_is_fun RS apply_type);
@@ -127,12 +129,12 @@
(*Every subset of a finite function is a finite function.*)
goal Finite.thy "!!b A. b: A-||>B ==> ALL z. z<=b --> z: A-||>B";
by (etac FiniteFun.induct 1);
-by (simp_tac (ZF_ss addsimps subset_empty_iff::FiniteFun.intrs) 1);
-by (asm_simp_tac (ZF_ss addsimps subset_cons_iff::distrib_simps) 1);
-by (safe_tac ZF_cs);
+by (simp_tac (!simpset addsimps subset_empty_iff::FiniteFun.intrs) 1);
+by (asm_simp_tac (!simpset addsimps subset_cons_iff::distrib_simps) 1);
+by (safe_tac (!claset));
by (eres_inst_tac [("b","z")] (cons_Diff RS subst) 1);
by (dtac (spec RS mp) 1 THEN assume_tac 1);
-by (fast_tac (ZF_cs addSIs FiniteFun.intrs) 1);
+by (fast_tac (!claset addSIs FiniteFun.intrs) 1);
qed "FiniteFun_subset_lemma";
goal Finite.thy "!!c b A. [| c<=b; b: A-||>B |] ==> c: A-||>B";
--- a/src/ZF/Fixedpt.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Fixedpt.ML Fri Jan 03 15:01:55 1997 +0100
@@ -56,13 +56,13 @@
(*lfp is contained in each pre-fixedpoint*)
goalw Fixedpt.thy [lfp_def]
"!!A. [| h(A) <= A; A<=D |] ==> lfp(D,h) <= A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
(*or by (rtac (PowI RS CollectI RS Inter_lower) 1); *)
qed "lfp_lowerbound";
(*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*)
goalw Fixedpt.thy [lfp_def,Inter_def] "lfp(D,h) <= D";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "lfp_subset";
(*Used in datatype package*)
@@ -190,7 +190,7 @@
qed "gfp_upperbound";
goalw Fixedpt.thy [gfp_def] "gfp(D,h) <= D";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "gfp_subset";
(*Used in datatype package*)
--- a/src/ZF/Fixedpt.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Fixedpt.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
Least and greatest fixed points
*)
-Fixedpt = ZF + "domrange" +
+Fixedpt = domrange +
consts
bnd_mono :: [i,i=>i]=>o
lfp, gfp :: [i,i=>i]=>i
--- a/src/ZF/IMP/Com.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/IMP/Com.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,11 +8,11 @@
val assign_type = prove_goalw Com.thy [assign_def]
"!!n. [| sigma:loc -> nat; n:nat |] ==> sigma[n/x] : loc -> nat"
- (fn _ => [ fast_tac (ZF_cs addIs [apply_type,lam_type,if_type]) 1 ]);
+ (fn _ => [ fast_tac (!claset addIs [apply_type,lam_type,if_type]) 1 ]);
-val type_cs = ZF_cs addSDs [evala.dom_subset RS subsetD,
- evalb.dom_subset RS subsetD,
- evalc.dom_subset RS subsetD];
+val type_cs = !claset addSDs [evala.dom_subset RS subsetD,
+ evalb.dom_subset RS subsetD,
+ evalc.dom_subset RS subsetD];
(********** type_intrs for evala **********)
--- a/src/ZF/IMP/Denotation.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/IMP/Denotation.ML Fri Jan 03 15:01:55 1997 +0100
@@ -23,7 +23,7 @@
"!!a.[|a:aexp; sigma:loc->nat|] ==> A(a,sigma):nat"
(fn _ => [(etac aexp.induct 1),
(rewrite_goals_tac A_rewrite_rules),
- (ALLGOALS (fast_tac (ZF_cs addSIs [apply_type])))]);
+ (ALLGOALS (fast_tac (!claset addSIs [apply_type])))]);
(**** Type_intr for B ****)
@@ -31,7 +31,7 @@
"!!b. [|b:bexp; sigma:loc->nat|] ==> B(b,sigma):bool"
(fn _ => [(etac bexp.induct 1),
(rewrite_goals_tac B_rewrite_rules),
- (ALLGOALS (fast_tac (ZF_cs
+ (ALLGOALS (fast_tac (!claset
addSIs [apply_type,A_type]@bool_typechecks)))]);
(**** C_subset ****)
@@ -40,7 +40,7 @@
"!!c. c:com ==> C(c) <= (loc->nat)*(loc->nat)"
(fn _ => [(etac com.induct 1),
(rewrite_tac C_rewrite_rules),
- (ALLGOALS (fast_tac (comp_cs addDs [lfp_subset RS subsetD])))]);
+ (ALLGOALS (fast_tac (!claset addDs [lfp_subset RS subsetD])))]);
(**** Type_elims for C ****)
@@ -49,7 +49,7 @@
\ !!c. [| x:loc->nat; y:loc->nat |] ==> R |] \
\ ==> R"
(fn prems => [(cut_facts_tac prems 1),
- (fast_tac (ZF_cs addSIs prems
+ (fast_tac (!claset addSIs prems
addDs [(C_subset RS subsetD)]) 1)]);
val C_type_fst = prove_goal Denotation.thy
@@ -61,13 +61,13 @@
(dtac (C_subset RS subsetD) 1),
(atac 1),
(etac SigmaE 1),
- (asm_simp_tac ZF_ss 1)]);
+ (Asm_simp_tac 1)]);
(**** bnd_mono (nat->nat*nat->nat,Gamma(b,c) ****)
val Gamma_bnd_mono = prove_goalw Denotation.thy [bnd_mono_def,Gamma_def]
"!!c. c:com ==> bnd_mono ((loc->nat)*(loc->nat),Gamma(b,c))"
- (fn prems => [(best_tac (comp_cs addEs [C_type]) 1)]);
+ (fn prems => [(best_tac (!claset addEs [C_type]) 1)]);
(**** End ***)
--- a/src/ZF/IMP/Equiv.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/IMP/Equiv.ML Fri Jan 03 15:01:55 1997 +0100
@@ -10,15 +10,15 @@
by (res_inst_tac [("x","a")] aexp.induct 1); (* struct. ind. *)
by (resolve_tac prems 1); (* type prem. *)
by (rewrite_goals_tac A_rewrite_rules); (* rewr. Den. *)
-by (TRYALL (fast_tac (ZF_cs addSIs (evala.intrs@prems)
- addSEs aexp_elim_cases)));
+by (TRYALL (fast_tac (!claset addSIs (evala.intrs@prems)
+ addSEs aexp_elim_cases)));
qed "aexp_iff";
val aexp1 = prove_goal Equiv.thy
"[| <a,sigma> -a-> n; a: aexp; sigma: loc -> nat |] \
\ ==> A(a,sigma) = n" (* destruction rule *)
- (fn prems => [(fast_tac (ZF_cs addSIs ((aexp_iff RS iffD1)::prems)) 1)]);
+ (fn prems => [(fast_tac (!claset addSIs ((aexp_iff RS iffD1)::prems)) 1)]);
val aexp2 = aexp_iff RS iffD2;
@@ -39,14 +39,14 @@
by (res_inst_tac [("x","b")] bexp.induct 1); (* struct. ind. *)
by (resolve_tac prems 1); (* type prem. *)
by (rewrite_goals_tac B_rewrite_rules); (* rewr. Den. *)
-by (TRYALL (fast_tac (ZF_cs addSIs (evalb.intrs@prems@[aexp2])
+by (TRYALL (fast_tac (!claset addSIs (evalb.intrs@prems@[aexp2])
addSDs [aexp1] addSEs bexp_elim_cases)));
qed "bexp_iff";
val bexp1 = prove_goal Equiv.thy
"[| <b,sigma> -b-> w; b: bexp; sigma: loc -> nat |]\
\ ==> B(b,sigma) = w"
- (fn prems => [(fast_tac (ZF_cs addSIs ((bexp_iff RS iffD1)::prems)) 1)]);
+ (fn prems => [(fast_tac (!claset addSIs ((bexp_iff RS iffD1)::prems)) 1)]);
val bexp2 = bexp_iff RS iffD2;
goal Equiv.thy "!!c. <c,sigma> -c-> sigma' ==> <sigma,sigma'> : C(c)";
@@ -56,80 +56,72 @@
by (rewrite_tac (Gamma_def::C_rewrite_rules));
(* skip *)
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
(* assign *)
-by (asm_full_simp_tac (ZF_ss addsimps [aexp1,assign_type] @ op_type_intrs) 1);
+by (asm_full_simp_tac (!simpset addsimps [aexp1,assign_type] @ op_type_intrs) 1);
(* comp *)
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
(* if *)
-by (asm_simp_tac (ZF_ss addsimps [bexp1]) 1);
-by (asm_simp_tac (ZF_ss addsimps [bexp1]) 1);
+by (asm_simp_tac (!simpset addsimps [bexp1]) 1);
+by (asm_simp_tac (!simpset addsimps [bexp1]) 1);
(* while *)
by (etac (rewrite_rule [Gamma_def]
(Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1);
-by (asm_simp_tac (ZF_ss addsimps [bexp1]) 1);
-by (fast_tac (comp_cs addSIs [bexp1,idI]@evalb_type_intrs) 1);
+by (asm_simp_tac (!simpset addsimps [bexp1]) 1);
+by (fast_tac (!claset addSIs [bexp1,idI]@evalb_type_intrs) 1);
by (etac (rewrite_rule [Gamma_def]
(Gamma_bnd_mono RS lfp_Tarski RS ssubst)) 1);
-by (asm_simp_tac (ZF_ss addsimps [bexp1]) 1);
-by (fast_tac (comp_cs addSIs [bexp1,compI]@evalb_type_intrs) 1);
+by (asm_simp_tac (!simpset addsimps [bexp1]) 1);
+by (fast_tac (!claset addSIs [bexp1,compI]@evalb_type_intrs) 1);
val com1 = result();
-val com_cs = ZF_cs addSIs [aexp2,bexp2,B_type,A_type]
- addIs evalc.intrs
- addSEs [idE,compE]
- addEs [C_type,C_type_fst];
+AddSIs [aexp2,bexp2,B_type,A_type];
+AddIs evalc.intrs;
+AddEs [C_type,C_type_fst];
val [prem] = goal Equiv.thy
"c : com ==> ALL x:C(c). <c,fst(x)> -c-> snd(x)";
by (rtac (prem RS com.induct) 1);
by (rewrite_tac C_rewrite_rules);
-by (safe_tac com_cs);
-by (ALLGOALS (asm_full_simp_tac ZF_ss));
+by (safe_tac (!claset));
+by (ALLGOALS Asm_full_simp_tac);
(* skip *)
-by (fast_tac com_cs 1);
+by (Fast_tac 1);
(* assign *)
-by (fast_tac com_cs 1);
+by (Fast_tac 1);
(* comp *)
by (REPEAT (EVERY [(dtac bspec 1),(atac 1)]));
-by (asm_full_simp_tac ZF_ss 1);
-by (fast_tac com_cs 1);
+by (Asm_full_simp_tac 1);
+by (Fast_tac 1);
(* while *)
by (EVERY1 [forward_tac [Gamma_bnd_mono], etac induct, atac]);
by (rewtac Gamma_def);
-by (safe_tac com_cs);
+by (safe_tac (!claset));
by (EVERY1 [dtac bspec, atac]);
-by (ALLGOALS (asm_full_simp_tac ZF_ss));
+by (ALLGOALS Asm_full_simp_tac);
(* while, if *)
-by (ALLGOALS (fast_tac com_cs));
+by (ALLGOALS Fast_tac);
val com2 = result();
(**** Proof of Equivalence ****)
-val com_iff_cs = ZF_cs addIs [C_subset RS subsetD]
- addEs [com2 RS bspec]
- addDs [com1];
-
goal Equiv.thy
"ALL c:com. C(c) = {io:(loc->nat)*(loc->nat). <c,fst(io)> -c-> snd(io)}";
-by (rtac ballI 1);
-by (rtac equalityI 1);
-(* => *)
-by (fast_tac com_iff_cs 1);
-(* <= *)
-by (REPEAT (step_tac com_iff_cs 1));
-by (asm_full_simp_tac ZF_ss 1);
+by (fast_tac (!claset addIs [equalityI, C_subset RS subsetD]
+ addEs [com2 RS bspec]
+ addDs [com1]
+ addss (!simpset)) 1);
val com_equivalence = result();
--- a/src/ZF/InfDatatype.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/InfDatatype.ML Fri Jan 03 15:01:55 1997 +0100
@@ -18,12 +18,12 @@
by (rtac le_UN_Ord_lt_csucc 2);
by (rtac ballI 4 THEN
etac fun_Limit_VfromE 4 THEN REPEAT_SOME assume_tac);
-by (fast_tac (ZF_cs addEs [Least_le RS lt_trans1, ltE]) 2);
+by (fast_tac (!claset addEs [Least_le RS lt_trans1, ltE]) 2);
by (rtac Pi_type 1);
by (rename_tac "w" 2);
by (etac fun_Limit_VfromE 2 THEN REPEAT_SOME assume_tac);
by (subgoal_tac "f`w : Vfrom(A, LEAST i. f`w : Vfrom(A,i))" 1);
-by (fast_tac (ZF_cs addEs [LeastI, ltE]) 2);
+by (fast_tac (!claset addEs [LeastI, ltE]) 2);
by (eresolve_tac [[subset_refl, UN_upper] MRS Vfrom_mono RS subsetD] 1);
by (assume_tac 1);
qed "fun_Vcsucc_lemma";
@@ -31,20 +31,20 @@
goal InfDatatype.thy
"!!K. [| W <= Vfrom(A,csucc(K)); |W| le K; InfCard(K) \
\ |] ==> EX j. W <= Vfrom(A,j) & j < csucc(K)";
-by (asm_full_simp_tac (ZF_ss addsimps [subset_iff_id, fun_Vcsucc_lemma]) 1);
+by (asm_full_simp_tac (!simpset addsimps [subset_iff_id, fun_Vcsucc_lemma]) 1);
qed "subset_Vcsucc";
(*Version for arbitrary index sets*)
goal InfDatatype.thy
"!!K. [| |W| le K; InfCard(K); W <= Vfrom(A,csucc(K)) |] ==> \
\ W -> Vfrom(A,csucc(K)) <= Vfrom(A,csucc(K))";
-by (safe_tac (ZF_cs addSDs [fun_Vcsucc_lemma, subset_Vcsucc]));
+by (safe_tac (!claset addSDs [fun_Vcsucc_lemma, subset_Vcsucc]));
by (resolve_tac [Vfrom RS ssubst] 1);
by (dtac fun_is_rel 1);
(*This level includes the function, and is below csucc(K)*)
by (res_inst_tac [("a1", "succ(succ(j Un ja))")] (UN_I RS UnI2) 1);
by (eresolve_tac [subset_trans RS PowI] 2);
-by (fast_tac (ZF_cs addIs [Pair_in_Vfrom, Vfrom_UnI1, Vfrom_UnI2]) 2);
+by (fast_tac (!claset addIs [Pair_in_Vfrom, Vfrom_UnI1, Vfrom_UnI2]) 2);
by (REPEAT (ares_tac [ltD, InfCard_csucc, InfCard_is_Limit,
Limit_has_succ, Un_least_lt] 1));
qed "fun_Vcsucc";
--- a/src/ZF/List.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/List.ML Fri Jan 03 15:01:55 1997 +0100
@@ -10,6 +10,9 @@
(*** Aspects of the datatype definition ***)
+Addsimps list.case_eqns;
+
+
(*An elimination rule, for type-checking*)
val ConsE = list.mk_cases list.con_defs "Cons(a,l) : list(A)";
@@ -42,7 +45,7 @@
goalw List.thy (list.defs@list.con_defs) "list(univ(A)) <= univ(A)";
by (rtac lfp_lowerbound 1);
by (rtac (A_subset_univ RS univ_mono) 2);
-by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
+by (fast_tac (!claset addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
Pair_in_univ]) 1);
qed "list_univ";
@@ -59,7 +62,7 @@
\ !!x y. [| x: A; y: list(A) |] ==> h(x,y): C(Cons(x,y)) \
\ |] ==> list_case(c,h,l) : C(l)";
by (rtac (major RS list.induct) 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps (list.case_eqns @ prems))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps (list.case_eqns @ prems))));
qed "list_case_type";
@@ -90,9 +93,11 @@
by (resolve_tac list.case_eqns 1);
qed "tl_Cons";
+Addsimps [hd_Cons, tl_Nil, tl_Cons];
+
goal List.thy "!!l. l: list(A) ==> tl(l) : list(A)";
by (etac list.elim 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps (list.intrs @ [tl_Nil,tl_Cons]))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps list.intrs)));
qed "tl_type";
(** drop **)
@@ -103,29 +108,30 @@
goalw List.thy [drop_def] "!!i. i:nat ==> drop(i, Nil) = Nil";
by (etac nat_induct 1);
-by (ALLGOALS (asm_simp_tac (nat_ss addsimps [tl_Nil])));
+by (ALLGOALS Asm_simp_tac);
qed "drop_Nil";
goalw List.thy [drop_def]
"!!i. i:nat ==> drop(succ(i), Cons(a,l)) = drop(i,l)";
+br sym 1;
by (etac nat_induct 1);
-by (simp_tac (nat_ss addsimps [tl_Cons]) 1);
-by (stac rec_succ 1);
-by (stac rec_succ 1);
-by (asm_simp_tac ZF_ss 1);
+by (Simp_tac 1);
+by (Asm_simp_tac 1);
qed "drop_succ_Cons";
+Addsimps [drop_0, drop_Nil, drop_succ_Cons];
+
goalw List.thy [drop_def]
"!!i l. [| i:nat; l: list(A) |] ==> drop(i,l) : list(A)";
by (etac nat_induct 1);
-by (ALLGOALS (asm_simp_tac (nat_ss addsimps [tl_type])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [tl_type])));
qed "drop_type";
(** list_rec -- by Vset recursion **)
goal List.thy "list_rec(Nil,c,h) = c";
by (rtac (list_rec_def RS def_Vrec RS trans) 1);
-by (simp_tac (ZF_ss addsimps list.case_eqns) 1);
+by (simp_tac (!simpset addsimps list.case_eqns) 1);
qed "list_rec_Nil";
goal List.thy "list_rec(Cons(a,l), c, h) = h(a, l, list_rec(l,c,h))";
@@ -133,6 +139,9 @@
by (simp_tac (rank_ss addsimps (rank_Cons2::list.case_eqns)) 1);
qed "list_rec_Cons";
+Addsimps [list_rec_Nil, list_rec_Cons];
+
+
(*Type checking -- proved by induction, as usual*)
val prems = goal List.thy
"[| l: list(A); \
@@ -140,8 +149,7 @@
\ !!x y r. [| x:A; y: list(A); r: C(y) |] ==> h(x,y,r): C(Cons(x,y)) \
\ |] ==> list_rec(l,c,h) : C(l)";
by (list_ind_tac "l" prems 1);
-by (ALLGOALS (asm_simp_tac
- (ZF_ss addsimps (prems@[list_rec_Nil,list_rec_Cons]))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "list_rec_type";
(** Versions for use with definitions **)
@@ -164,6 +172,7 @@
(** map **)
val [map_Nil,map_Cons] = list_recs map_def;
+Addsimps [map_Nil, map_Cons];
val prems = goalw List.thy [map_def]
"[| l: list(A); !!x. x: A ==> h(x): B |] ==> map(h,l) : list(B)";
@@ -178,6 +187,7 @@
(** length **)
val [length_Nil,length_Cons] = list_recs length_def;
+Addsimps [length_Nil,length_Cons];
goalw List.thy [length_def]
"!!l. l: list(A) ==> length(l) : nat";
@@ -187,6 +197,7 @@
(** app **)
val [app_Nil,app_Cons] = list_recs app_def;
+Addsimps [app_Nil, app_Cons];
goalw List.thy [app_def]
"!!xs ys. [| xs: list(A); ys: list(A) |] ==> xs@ys : list(A)";
@@ -196,6 +207,7 @@
(** rev **)
val [rev_Nil,rev_Cons] = list_recs rev_def;
+Addsimps [rev_Nil,rev_Cons] ;
goalw List.thy [rev_def]
"!!xs. xs: list(A) ==> rev(xs) : list(A)";
@@ -206,6 +218,7 @@
(** flat **)
val [flat_Nil,flat_Cons] = list_recs flat_def;
+Addsimps [flat_Nil,flat_Cons];
goalw List.thy [flat_def]
"!!ls. ls: list(list(A)) ==> flat(ls) : list(A)";
@@ -216,45 +229,38 @@
(** set_of_list **)
val [set_of_list_Nil,set_of_list_Cons] = list_recs set_of_list_def;
+Addsimps [set_of_list_Nil,set_of_list_Cons];
goalw List.thy [set_of_list_def]
"!!l. l: list(A) ==> set_of_list(l) : Pow(A)";
by (etac list_rec_type 1);
-by (ALLGOALS (fast_tac ZF_cs));
+by (ALLGOALS (Fast_tac));
qed "set_of_list_type";
goal List.thy
"!!l. xs: list(A) ==> \
\ set_of_list (xs@ys) = set_of_list(xs) Un set_of_list(ys)";
by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [set_of_list_Nil,set_of_list_Cons,
- app_Nil,app_Cons,Un_cons])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [Un_cons])));
qed "set_of_list_append";
(** list_add **)
val [list_add_Nil,list_add_Cons] = list_recs list_add_def;
+Addsimps [list_add_Nil,list_add_Cons];
goalw List.thy [list_add_def]
"!!xs. xs: list(nat) ==> list_add(xs) : nat";
by (REPEAT (ares_tac [list_rec_type, nat_0I, add_type] 1));
qed "list_add_type";
-(** List simplification **)
-
val list_typechecks =
list.intrs @
[list_rec_type, map_type, map_type2, app_type, length_type,
rev_type, flat_type, list_add_type];
-val list_ss = arith_ss
- addsimps list.case_eqns
- addsimps [list_rec_Nil, list_rec_Cons,
- map_Nil, map_Cons, app_Nil, app_Cons,
- length_Nil, length_Cons, rev_Nil, rev_Cons,
- flat_Nil, flat_Cons, list_add_Nil, list_add_Cons]
- setsolver (type_auto_tac list_typechecks);
+simpset := !simpset setsolver (type_auto_tac list_typechecks);
(*** theorems about map ***)
@@ -262,25 +268,25 @@
val prems = goal List.thy
"l: list(A) ==> map(%u.u, l) = l";
by (list_ind_tac "l" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "map_ident";
val prems = goal List.thy
"l: list(A) ==> map(h, map(j,l)) = map(%u.h(j(u)), l)";
by (list_ind_tac "l" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "map_compose";
val prems = goal List.thy
"xs: list(A) ==> map(h, xs@ys) = map(h,xs) @ map(h,ys)";
by (list_ind_tac "xs" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "map_app_distrib";
val prems = goal List.thy
"ls: list(list(A)) ==> map(h, flat(ls)) = flat(map(map(h),ls))";
by (list_ind_tac "ls" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_app_distrib])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [map_app_distrib])));
qed "map_flat";
val prems = goal List.thy
@@ -288,7 +294,7 @@
\ list_rec(map(h,l), c, d) = \
\ list_rec(l, c, %x xs r. d(h(x), map(h,xs), r))";
by (list_ind_tac "l" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "list_rec_map";
(** theorems about list(Collect(A,P)) -- used in ex/term.ML **)
@@ -299,7 +305,7 @@
val prems = goal List.thy
"l: list({x:A. h(x)=j(x)}) ==> map(h,l) = map(j,l)";
by (list_ind_tac "l" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "map_list_Collect";
(*** theorems about length ***)
@@ -307,13 +313,13 @@
val prems = goal List.thy
"xs: list(A) ==> length(map(h,xs)) = length(xs)";
by (list_ind_tac "xs" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "length_map";
val prems = goal List.thy
"xs: list(A) ==> length(xs@ys) = length(xs) #+ length(ys)";
by (list_ind_tac "xs" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "length_app";
(* [| m: nat; n: nat |] ==> m #+ succ(n) = succ(n) #+ m
@@ -323,13 +329,13 @@
val prems = goal List.thy
"xs: list(A) ==> length(rev(xs)) = length(xs)";
by (list_ind_tac "xs" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [length_app, add_commute_succ])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [length_app, add_commute_succ])));
qed "length_rev";
val prems = goal List.thy
"ls: list(list(A)) ==> length(flat(ls)) = list_add(map(length,ls))";
by (list_ind_tac "ls" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [length_app])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [length_app])));
qed "length_flat";
(** Length and drop **)
@@ -339,27 +345,27 @@
"!!xs. xs: list(A) ==> \
\ ALL x. EX z zs. drop(length(xs), Cons(x,xs)) = Cons(z,zs)";
by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [drop_0,drop_succ_Cons])));
-by (fast_tac ZF_cs 1);
+by (ALLGOALS Asm_simp_tac);
+by (Fast_tac 1);
qed "drop_length_Cons_lemma";
bind_thm ("drop_length_Cons", (drop_length_Cons_lemma RS spec));
goal List.thy
"!!l. l: list(A) ==> ALL i: length(l). EX z zs. drop(i,l) = Cons(z,zs)";
by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps bquant_simps)));
+by (ALLGOALS Asm_simp_tac);
by (rtac conjI 1);
by (etac drop_length_Cons 1);
by (rtac ballI 1);
by (rtac natE 1);
by (etac ([asm_rl, length_type, Ord_nat] MRS Ord_trans) 1);
by (assume_tac 1);
-by (asm_simp_tac (list_ss addsimps [drop_0]) 1);
-by (fast_tac ZF_cs 1);
-by (asm_simp_tac (list_ss addsimps [drop_succ_Cons]) 1);
+by (Asm_simp_tac 1);
+by (Fast_tac 1);
+by (Asm_simp_tac 1);
by (dtac bspec 1);
-by (fast_tac ZF_cs 2);
-by (fast_tac (ZF_cs addEs [succ_in_naturalD,length_type]) 1);
+by (Fast_tac 2);
+by (fast_tac (!claset addEs [succ_in_naturalD,length_type]) 1);
qed "drop_length_lemma";
bind_thm ("drop_length", (drop_length_lemma RS bspec));
@@ -368,25 +374,25 @@
val [major] = goal List.thy "xs: list(A) ==> xs@Nil=xs";
by (rtac (major RS list.induct) 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "app_right_Nil";
val prems = goal List.thy "xs: list(A) ==> (xs@ys)@zs = xs@(ys@zs)";
by (list_ind_tac "xs" prems 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "app_assoc";
val prems = goal List.thy
"ls: list(list(A)) ==> flat(ls@ms) = flat(ls)@flat(ms)";
by (list_ind_tac "ls" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [app_assoc])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [app_assoc])));
qed "flat_app_distrib";
(*** theorems about rev ***)
val prems = goal List.thy "l: list(A) ==> rev(map(h,l)) = map(h,rev(l))";
by (list_ind_tac "l" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [map_app_distrib])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [map_app_distrib])));
qed "rev_map_distrib";
(*Simplifier needs the premises as assumptions because rewriting will not
@@ -396,18 +402,18 @@
goal List.thy
"!!xs. [| xs: list(A); ys: list(A) |] ==> rev(xs@ys) = rev(ys)@rev(xs)";
by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [app_right_Nil,app_assoc])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [app_right_Nil,app_assoc])));
qed "rev_app_distrib";
val prems = goal List.thy "l: list(A) ==> rev(rev(l))=l";
by (list_ind_tac "l" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [rev_app_distrib])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [rev_app_distrib])));
qed "rev_rev_ident";
val prems = goal List.thy
"ls: list(list(A)) ==> rev(flat(ls)) = flat(map(rev,rev(ls)))";
by (list_ind_tac "ls" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps
+by (ALLGOALS (asm_simp_tac (!simpset addsimps
[map_app_distrib, flat_app_distrib, rev_app_distrib, app_right_Nil])));
qed "rev_flat";
@@ -420,7 +426,7 @@
by (cut_facts_tac prems 1);
by (list_ind_tac "xs" prems 1);
by (ALLGOALS
- (asm_simp_tac (list_ss addsimps [add_0_right, add_assoc RS sym])));
+ (asm_simp_tac (!simpset addsimps [add_0_right, add_assoc RS sym])));
by (rtac (add_commute RS subst_context) 1);
by (REPEAT (ares_tac [refl, list_add_type] 1));
qed "list_add_app";
@@ -429,13 +435,13 @@
"l: list(nat) ==> list_add(rev(l)) = list_add(l)";
by (list_ind_tac "l" prems 1);
by (ALLGOALS
- (asm_simp_tac (list_ss addsimps [list_add_app, add_0_right])));
+ (asm_simp_tac (!simpset addsimps [list_add_app, add_0_right])));
qed "list_add_rev";
val prems = goal List.thy
"ls: list(list(nat)) ==> list_add(flat(ls)) = list_add(map(list_add,ls))";
by (list_ind_tac "ls" prems 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [list_add_app])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [list_add_app])));
by (REPEAT (ares_tac [refl, list_add_type, map_type, add_commute] 1));
qed "list_add_flat";
@@ -448,6 +454,6 @@
\ |] ==> P(l)";
by (rtac (major RS rev_rev_ident RS subst) 1);
by (rtac (major RS rev_type RS list.induct) 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps prems)));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "list_append_induct";
--- a/src/ZF/Makefile Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Makefile Fri Jan 03 15:01:55 1997 +0100
@@ -22,7 +22,7 @@
BIN = $(ISABELLEBIN)
COMP = $(ISABELLECOMP)
NAMES = ZF upair subset pair domrange \
- func AC simpdata equalities Bool \
+ func AC equalities Bool \
Sum QPair mono Fixedpt ind_syntax cartprod add_ind_def \
constructor intr_elim indrule Inductive Perm Rel EquivClass Trancl \
WF Order Ordinal Epsilon Arith Univ \
@@ -30,7 +30,7 @@
Cardinal CardinalArith Cardinal_AC InfDatatype \
Zorn Nat Finite List
-FILES = ROOT.ML thy_syntax.ML ../Pure/section_utils.ML \
+FILES = ROOT.ML thy_syntax.ML ../Pure/section_utils.ML simpdata.ML typechk.ML\
$(NAMES:%=%.thy) $(NAMES:%=%.ML)
#Uses cp rather than make_database because Poly/ML allows only 3 levels
@@ -83,9 +83,9 @@
fi
##Coinduction example
-COIND_NAMES = ECR Language MT Map Static Types Values
+COIND_NAMES = ECR MT Map Static Types Values
-COIND_FILES = Coind/ROOT.ML Coind/BCR.thy Coind/Dynamic.thy \
+COIND_FILES = Coind/ROOT.ML Coind/Language.thy Coind/BCR.thy Coind/Dynamic.thy\
$(COIND_NAMES:%=Coind/%.thy) $(COIND_NAMES:%=Coind/%.ML)
Coind: $(BIN)/ZF $(COIND_FILES)
@@ -95,10 +95,10 @@
fi
##AC examples
-AC_NAMES = AC_Equiv OrdQuant Transrec2 Cardinal_aux \
+AC_NAMES = AC_Equiv Cardinal_aux \
AC15_WO6 AC16_WO4 AC16_lemmas AC17_AC1 AC18_AC19 AC1_WO2 \
DC DC_lemmas HH Hartog WO1_AC \
- WO2_AC16 WO6_WO1 WO_AC first recfunAC16 rel_is_fun
+ WO2_AC16 WO6_WO1 WO_AC recfunAC16 rel_is_fun
AC_FILES = AC/ROOT.ML AC/AC0_AC1.ML AC/AC10_AC15.ML AC/AC1_AC17.ML \
AC/AC2_AC6.ML AC/AC7_AC9.ML \
--- a/src/ZF/Nat.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Nat.ML Fri Jan 03 15:01:55 1997 +0100
@@ -12,7 +12,7 @@
by (rtac bnd_monoI 1);
by (REPEAT (ares_tac [subset_refl, RepFun_mono, Un_mono] 2));
by (cut_facts_tac [infinity] 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "nat_bnd_mono";
(* nat = {0} Un {succ(x). x:nat} *)
@@ -39,6 +39,9 @@
by (rtac (nat_1I RS nat_succI) 1);
qed "nat_2I";
+Addsimps [nat_0I, nat_1I, nat_2I, nat_succI];
+AddSIs [nat_0I, nat_1I, nat_2I];
+
goal Nat.thy "bool <= nat";
by (REPEAT (ares_tac [subsetI,nat_0I,nat_1I] 1
ORELSE eresolve_tac [boolE,ssubst] 1));
@@ -53,7 +56,7 @@
val major::prems = goal Nat.thy
"[| n: nat; P(0); !!x. [| x: nat; P(x) |] ==> P(succ(x)) |] ==> P(n)";
by (rtac ([nat_def, nat_bnd_mono, major] MRS def_induct) 1);
-by (fast_tac (ZF_cs addIs prems) 1);
+by (fast_tac (!claset addIs prems) 1);
qed "nat_induct";
(*Perform induction on n, then prove the n:nat subgoal using prems. *)
@@ -90,15 +93,18 @@
qed "Ord_nat";
goalw Nat.thy [Limit_def] "Limit(nat)";
-by (safe_tac (ZF_cs addSIs [ltI, nat_0I, nat_1I, nat_succI, Ord_nat]));
+by (safe_tac (!claset addSIs [ltI, nat_succI, Ord_nat]));
by (etac ltD 1);
qed "Limit_nat";
+Addsimps [Ord_nat, Limit_nat];
+AddSIs [Ord_nat, Limit_nat];
+
goal Nat.thy "!!i. Limit(i) ==> nat le i";
by (rtac subset_imp_le 1);
by (rtac subsetI 1);
by (etac nat_induct 1);
-by (fast_tac (ZF_cs addIs [Limit_has_succ RS ltD, ltI, Limit_is_Ord]) 2);
+by (fast_tac (!claset addIs [Limit_has_succ RS ltD, ltI, Limit_is_Ord]) 2);
by (REPEAT (ares_tac [Limit_has_0 RS ltD,
Ord_nat, Limit_is_Ord] 1));
qed "nat_le_Limit";
@@ -128,7 +134,7 @@
by (nat_ind_tac "n" prems 1);
by (ALLGOALS
(asm_simp_tac
- (ZF_ss addsimps (prems@distrib_simps@[le0_iff, le_succ_iff]))));
+ (!simpset addsimps (prems@distrib_simps@[le0_iff, le_succ_iff]))));
qed "nat_induct_from_lemma";
(*Induction starting from m rather than 0*)
@@ -164,7 +170,7 @@
by (etac nat_induct 1);
by (ALLGOALS
(EVERY' [rtac (impI RS impI), rtac (nat_induct RS ballI), assume_tac,
- fast_tac lt_cs, fast_tac lt_cs]));
+ fast_tac le_cs, fast_tac le_cs]));
qed "succ_lt_induct_lemma";
val prems = goal Nat.thy
@@ -180,19 +186,20 @@
(** nat_case **)
goalw Nat.thy [nat_case_def] "nat_case(a,b,0) = a";
-by (fast_tac (ZF_cs addIs [the_equality]) 1);
+by (fast_tac (!claset addIs [the_equality]) 1);
qed "nat_case_0";
goalw Nat.thy [nat_case_def] "nat_case(a,b,succ(m)) = b(m)";
-by (fast_tac (ZF_cs addIs [the_equality]) 1);
+by (fast_tac (!claset addIs [the_equality]) 1);
qed "nat_case_succ";
+Addsimps [nat_case_0, nat_case_succ];
+
val major::prems = goal Nat.thy
"[| n: nat; a: C(0); !!m. m: nat ==> b(m): C(succ(m)) \
\ |] ==> nat_case(a,b,n) : C(n)";
by (rtac (major RS nat_induct) 1);
-by (ALLGOALS
- (asm_simp_tac (ZF_ss addsimps (prems @ [nat_case_0, nat_case_succ]))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "nat_case_type";
@@ -209,8 +216,7 @@
val [prem] = goal Nat.thy
"m: nat ==> nat_rec(succ(m),a,b) = b(m, nat_rec(m,a,b))";
by (rtac nat_rec_trans 1);
-by (simp_tac (ZF_ss addsimps [prem, nat_case_succ, nat_succI, Memrel_iff,
- vimage_singleton_iff]) 1);
+by (simp_tac (!simpset addsimps [prem, Memrel_iff, vimage_singleton_iff]) 1);
qed "nat_rec_succ";
(** The union of two natural numbers is a natural number -- their maximum **)
--- a/src/ZF/Nat.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Nat.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
Natural numbers in Zermelo-Fraenkel Set Theory
*)
-Nat = Ordinal + Bool + "mono" +
+Nat = OrdQuant + Bool + mono +
consts
nat :: i
nat_case :: [i, i=>i, i]=>i
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/ZF/OrdQuant.ML Fri Jan 03 15:01:55 1997 +0100
@@ -0,0 +1,112 @@
+(* Title: ZF/AC/OrdQuant.thy
+ ID: $Id$
+ Authors: Krzysztof Grabczewski and L C Paulson
+
+Quantifiers and union operator for ordinals.
+*)
+
+open OrdQuant;
+
+(*** universal quantifier for ordinals ***)
+
+qed_goalw "oallI" thy [oall_def]
+ "[| !!x. x<A ==> P(x) |] ==> ALL x<A. P(x)"
+ (fn prems=> [ (REPEAT (ares_tac (prems @ [allI,impI]) 1)) ]);
+
+qed_goalw "ospec" thy [oall_def]
+ "[| ALL x<A. P(x); x<A |] ==> P(x)"
+ (fn major::prems=>
+ [ (rtac (major RS spec RS mp) 1),
+ (resolve_tac prems 1) ]);
+
+qed_goalw "oallE" thy [oall_def]
+ "[| ALL x<A. P(x); P(x) ==> Q; ~x<A ==> Q |] ==> Q"
+ (fn major::prems=>
+ [ (rtac (major RS allE) 1),
+ (REPEAT (eresolve_tac (prems@[asm_rl,impCE]) 1)) ]);
+
+qed_goal "rev_oallE" thy
+ "[| ALL x<A. P(x); ~x<A ==> Q; P(x) ==> Q |] ==> Q"
+ (fn major::prems=>
+ [ (rtac (major RS oallE) 1),
+ (REPEAT (eresolve_tac prems 1)) ]);
+
+(*Trival rewrite rule; (ALL x<a.P)<->P holds only if a is not 0!*)
+qed_goal "oall_simp" thy "(ALL x<a. True) <-> True"
+ (fn _=> [ (REPEAT (ares_tac [TrueI,oallI,iffI] 1)) ]);
+
+(*Congruence rule for rewriting*)
+qed_goalw "oall_cong" thy [oall_def]
+ "[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) |] ==> oall(a,P) <-> oall(a',P')"
+ (fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);
+
+
+(*** existential quantifier for ordinals ***)
+
+qed_goalw "oexI" thy [oex_def]
+ "[| P(x); x<A |] ==> EX x<A. P(x)"
+ (fn prems=> [ (REPEAT (ares_tac (prems @ [exI,conjI]) 1)) ]);
+
+(*Not of the general form for such rules; ~EX has become ALL~ *)
+qed_goal "oexCI" thy
+ "[| ALL x<A. ~P(x) ==> P(a); a<A |] ==> EX x<A.P(x)"
+ (fn prems=>
+ [ (rtac classical 1),
+ (REPEAT (ares_tac (prems@[oexI,oallI,notI,notE]) 1)) ]);
+
+qed_goalw "oexE" thy [oex_def]
+ "[| EX x<A. P(x); !!x. [| x<A; P(x) |] ==> Q \
+\ |] ==> Q"
+ (fn major::prems=>
+ [ (rtac (major RS exE) 1),
+ (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]);
+
+qed_goalw "oex_cong" thy [oex_def]
+ "[| a=a'; !!x. x<a' ==> P(x) <-> P'(x) \
+\ |] ==> oex(a,P) <-> oex(a',P')"
+ (fn prems=> [ (simp_tac (!simpset addsimps prems addcongs [conj_cong]) 1) ]);
+
+
+(*** Rules for Ordinal-Indexed Unions ***)
+
+qed_goalw "OUN_I" thy [OUnion_def]
+ "!!i. [| a<i; b: B(a) |] ==> b: (UN z<i. B(z))"
+ (fn _=> [ fast_tac (!claset addSEs [ltE]) 1 ]);
+
+qed_goalw "OUN_E" thy [OUnion_def]
+ "[| b : (UN z<i. B(z)); !!a.[| b: B(a); a<i |] ==> R |] ==> R"
+ (fn major::prems=>
+ [ (rtac (major RS CollectE) 1),
+ (rtac UN_E 1),
+ (REPEAT (ares_tac (ltI::prems) 1)) ]);
+
+qed_goalw "OUN_iff" thy [oex_def]
+ "b : (UN x<i. B(x)) <-> (EX x<i. b : B(x))"
+ (fn _=> [ (fast_tac (!claset addIs [OUN_I] addSEs [OUN_E]) 1) ]);
+
+qed_goal "OUN_cong" thy
+ "[| i=j; !!x. x<j ==> C(x)=D(x) |] ==> (UN x<i.C(x)) = (UN x<j.D(x))"
+ (fn prems=>
+ [ rtac equality_iffI 1,
+ simp_tac (!simpset addcongs [oex_cong] addsimps (OUN_iff::prems)) 1 ]);
+
+AddSIs [oallI];
+AddIs [oexI, OUN_I];
+AddSEs [oexE, OUN_E];
+AddEs [rev_oallE];
+
+val Ord_atomize = atomize (("oall", [ospec])::ZF_conn_pairs,
+ ZF_mem_pairs);
+
+simpset := !simpset setmksimps (map mk_meta_eq o Ord_atomize o gen_all)
+ addsimps [oall_simp, ltD RS beta]
+ addcongs [oall_cong, oex_cong, OUN_cong];
+
+val major::prems = goalw thy [lt_def, oall_def]
+ "[| i<k; !!x.[| x<k; ALL y<x. P(y) |] ==> P(x) \
+\ |] ==> P(i)";
+by (rtac (major RS conjE) 1);
+by (etac Ord_induct 1 THEN assume_tac 1);
+by (fast_tac (!claset addIs prems) 1);
+qed "lt_induct";
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/ZF/OrdQuant.thy Fri Jan 03 15:01:55 1997 +0100
@@ -0,0 +1,38 @@
+(* Title: ZF/AC/OrdQuant.thy
+ ID: $Id$
+ Authors: Krzysztof Grabczewski and L C Paulson
+
+Quantifiers and union operator for ordinals.
+*)
+
+OrdQuant = Ordinal +
+
+consts
+
+ (* Ordinal Quantifiers *)
+ oall, oex :: [i, i => o] => o
+
+ (* Ordinal Union *)
+ OUnion :: [i, i => i] => i
+
+syntax
+
+ "@oall" :: [idt, i, o] => o ("(3ALL _<_./ _)" 10)
+ "@oex" :: [idt, i, o] => o ("(3EX _<_./ _)" 10)
+ "@OUNION" :: [idt, i, i] => i ("(3UN _<_./ _)" 10)
+
+translations
+
+ "ALL x<a. P" == "oall(a, %x. P)"
+ "EX x<a. P" == "oex(a, %x. P)"
+ "UN x<a. B" == "OUnion(a, %x. B)"
+
+defs
+
+ (* Ordinal Quantifiers *)
+ oall_def "oall(A, P) == ALL x. x<A --> P(x)"
+ oex_def "oex(A, P) == EX x. x<A & P(x)"
+
+ OUnion_def "OUnion(i,B) == {z: UN x:i. B(x). Ord(i)}"
+
+end
--- a/src/ZF/Order.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Order.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,15 +11,12 @@
open Order;
-val bij_apply_cs = ZF_cs addSEs [bij_converse_bij]
- addIs [bij_is_fun, apply_type];
-
(** Basic properties of the definitions **)
(*needed?*)
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, asym_def]
"!!r. part_ord(A,r) ==> asym(r Int A*A)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "part_ord_Imp_asym";
val major::premx::premy::prems = goalw Order.thy [linear_def]
@@ -41,30 +38,30 @@
goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def,
trans_on_def, well_ord_def]
"!!r. [| wf[A](r); linear(A,r) |] ==> well_ord(A,r)";
-by (asm_simp_tac (ZF_ss addsimps [wf_on_not_refl]) 1);
-by (fast_tac (ZF_cs addEs [linearE, wf_on_asym, wf_on_chain3]) 1);
+by (asm_simp_tac (!simpset addsimps [wf_on_not_refl]) 1);
+by (fast_tac (!claset addEs [linearE, wf_on_asym, wf_on_chain3]) 1);
qed "well_ordI";
goalw Order.thy [well_ord_def]
"!!r. well_ord(A,r) ==> wf[A](r)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
qed "well_ord_is_wf";
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def]
"!!r. well_ord(A,r) ==> trans[A](r)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
qed "well_ord_is_trans_on";
goalw Order.thy [well_ord_def, tot_ord_def]
"!!r. well_ord(A,r) ==> linear(A,r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "well_ord_is_linear";
(** Derived rules for pred(A,x,r) **)
goalw Order.thy [pred_def] "y : pred(A,x,r) <-> <y,x>:r & y:A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "pred_iff";
bind_thm ("predI", conjI RS (pred_iff RS iffD2));
@@ -76,11 +73,11 @@
qed "predE";
goalw Order.thy [pred_def] "pred(A,x,r) <= r -`` {x}";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "pred_subset_under";
goalw Order.thy [pred_def] "pred(A,x,r) <= A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "pred_subset";
goalw Order.thy [pred_def]
@@ -101,76 +98,76 @@
(*Note: a relation s such that s<=r need not be a partial ordering*)
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def]
"!!A B r. [| part_ord(A,r); B<=A |] ==> part_ord(B,r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "part_ord_subset";
goalw Order.thy [linear_def]
"!!A B r. [| linear(A,r); B<=A |] ==> linear(B,r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "linear_subset";
goalw Order.thy [tot_ord_def]
"!!A B r. [| tot_ord(A,r); B<=A |] ==> tot_ord(B,r)";
-by (fast_tac (ZF_cs addSEs [part_ord_subset, linear_subset]) 1);
+by (fast_tac (!claset addSEs [part_ord_subset, linear_subset]) 1);
qed "tot_ord_subset";
goalw Order.thy [well_ord_def]
"!!A B r. [| well_ord(A,r); B<=A |] ==> well_ord(B,r)";
-by (fast_tac (ZF_cs addSEs [tot_ord_subset, wf_on_subset_A]) 1);
+by (fast_tac (!claset addSEs [tot_ord_subset, wf_on_subset_A]) 1);
qed "well_ord_subset";
(** Relations restricted to a smaller domain, by Krzysztof Grabczewski **)
goalw Order.thy [irrefl_def] "irrefl(A,r Int A*A) <-> irrefl(A,r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "irrefl_Int_iff";
goalw Order.thy [trans_on_def] "trans[A](r Int A*A) <-> trans[A](r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "trans_on_Int_iff";
goalw Order.thy [part_ord_def] "part_ord(A,r Int A*A) <-> part_ord(A,r)";
-by (simp_tac (ZF_ss addsimps [irrefl_Int_iff, trans_on_Int_iff]) 1);
+by (simp_tac (!simpset addsimps [irrefl_Int_iff, trans_on_Int_iff]) 1);
qed "part_ord_Int_iff";
goalw Order.thy [linear_def] "linear(A,r Int A*A) <-> linear(A,r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "linear_Int_iff";
goalw Order.thy [tot_ord_def] "tot_ord(A,r Int A*A) <-> tot_ord(A,r)";
-by (simp_tac (ZF_ss addsimps [part_ord_Int_iff, linear_Int_iff]) 1);
+by (simp_tac (!simpset addsimps [part_ord_Int_iff, linear_Int_iff]) 1);
qed "tot_ord_Int_iff";
goalw Order.thy [wf_on_def, wf_def] "wf[A](r Int A*A) <-> wf[A](r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_on_Int_iff";
goalw Order.thy [well_ord_def] "well_ord(A,r Int A*A) <-> well_ord(A,r)";
-by (simp_tac (ZF_ss addsimps [tot_ord_Int_iff, wf_on_Int_iff]) 1);
+by (simp_tac (!simpset addsimps [tot_ord_Int_iff, wf_on_Int_iff]) 1);
qed "well_ord_Int_iff";
(** Relations over the Empty Set **)
goalw Order.thy [irrefl_def] "irrefl(0,r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "irrefl_0";
goalw Order.thy [trans_on_def] "trans[0](r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "trans_on_0";
goalw Order.thy [part_ord_def] "part_ord(0,r)";
-by (simp_tac (ZF_ss addsimps [irrefl_0, trans_on_0]) 1);
+by (simp_tac (!simpset addsimps [irrefl_0, trans_on_0]) 1);
qed "part_ord_0";
goalw Order.thy [linear_def] "linear(0,r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "linear_0";
goalw Order.thy [tot_ord_def] "tot_ord(0,r)";
-by (simp_tac (ZF_ss addsimps [part_ord_0, linear_0]) 1);
+by (simp_tac (!simpset addsimps [part_ord_0, linear_0]) 1);
qed "tot_ord_0";
goalw Order.thy [wf_on_def, wf_def] "wf[0](r)";
@@ -178,7 +175,7 @@
qed "wf_on_0";
goalw Order.thy [well_ord_def] "well_ord(0,r)";
-by (simp_tac (ZF_ss addsimps [tot_ord_0, wf_on_0]) 1);
+by (simp_tac (!simpset addsimps [tot_ord_0, wf_on_0]) 1);
qed "well_ord_0";
@@ -191,12 +188,12 @@
goalw Order.thy [mono_map_def, inj_def]
"!!f. [| linear(A,r); wf[B](s); f: mono_map(A,r,B,s) |] ==> f: inj(A,B)";
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
by (linear_case_tac 1);
by (REPEAT
(EVERY [eresolve_tac [wf_on_not_refl RS notE] 1,
etac ssubst 2,
- fast_tac ZF_cs 2,
+ Fast_tac 2,
REPEAT (ares_tac [apply_type] 1)]));
qed "mono_map_is_inj";
@@ -207,12 +204,12 @@
"[| f: bij(A, B); \
\ !!x y. [| x:A; y:A |] ==> <x, y> : r <-> <f`x, f`y> : s \
\ |] ==> f: ord_iso(A,r,B,s)";
-by (fast_tac (ZF_cs addSIs prems) 1);
+by (fast_tac (!claset addSIs prems) 1);
qed "ord_isoI";
goalw Order.thy [ord_iso_def, mono_map_def]
"!!f. f: ord_iso(A,r,B,s) ==> f: mono_map(A,r,B,s)";
-by (fast_tac (ZF_cs addSDs [bij_is_fun]) 1);
+by (fast_tac (!claset addSDs [bij_is_fun]) 1);
qed "ord_iso_is_mono_map";
goalw Order.thy [ord_iso_def]
@@ -224,7 +221,7 @@
goalw Order.thy [ord_iso_def]
"!!f. [| f: ord_iso(A,r,B,s); <x,y>: r; x:A; y:A |] ==> \
\ <f`x, f`y> : s";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "ord_iso_apply";
goalw Order.thy [ord_iso_def]
@@ -234,15 +231,15 @@
by (etac (bspec RS bspec RS iffD2) 1);
by (REPEAT (eresolve_tac [asm_rl,
bij_converse_bij RS bij_is_fun RS apply_type] 1));
-by (asm_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1);
+by (asm_simp_tac (!simpset addsimps [right_inverse_bij]) 1);
qed "ord_iso_converse";
(*Rewriting with bijections and converse (function inverse)*)
val bij_inverse_ss =
- ZF_ss setsolver (type_auto_tac [ord_iso_is_bij, bij_is_fun, apply_type,
- bij_converse_bij, comp_fun, comp_bij])
- addsimps [right_inverse_bij, left_inverse_bij, comp_fun_apply];
+ !simpset setsolver (type_auto_tac [ord_iso_is_bij, bij_is_fun, apply_type,
+ bij_converse_bij, comp_fun, comp_bij])
+ addsimps [right_inverse_bij, left_inverse_bij];
(** Symmetry and Transitivity Rules **)
@@ -250,27 +247,27 @@
(*Reflexivity of similarity*)
goal Order.thy "id(A): ord_iso(A,r,A,r)";
by (resolve_tac [id_bij RS ord_isoI] 1);
-by (asm_simp_tac (ZF_ss addsimps [id_conv]) 1);
+by (Asm_simp_tac 1);
qed "ord_iso_refl";
(*Symmetry of similarity*)
goalw Order.thy [ord_iso_def]
"!!f. f: ord_iso(A,r,B,s) ==> converse(f): ord_iso(B,s,A,r)";
-by (fast_tac (ZF_cs addss bij_inverse_ss) 1);
+by (fast_tac (!claset addss bij_inverse_ss) 1);
qed "ord_iso_sym";
(*Transitivity of similarity*)
goalw Order.thy [mono_map_def]
"!!f. [| g: mono_map(A,r,B,s); f: mono_map(B,s,C,t) |] ==> \
\ (f O g): mono_map(A,r,C,t)";
-by (fast_tac (ZF_cs addss bij_inverse_ss) 1);
+by (fast_tac (!claset addss bij_inverse_ss) 1);
qed "mono_map_trans";
(*Transitivity of similarity: the order-isomorphism relation*)
goalw Order.thy [ord_iso_def]
"!!f. [| g: ord_iso(A,r,B,s); f: ord_iso(B,s,C,t) |] ==> \
\ (f O g): ord_iso(A,r,C,t)";
-by (fast_tac (ZF_cs addss bij_inverse_ss) 1);
+by (fast_tac (!claset addss bij_inverse_ss) 1);
qed "ord_iso_trans";
(** Two monotone maps can make an order-isomorphism **)
@@ -278,13 +275,12 @@
goalw Order.thy [ord_iso_def, mono_map_def]
"!!f g. [| f: mono_map(A,r,B,s); g: mono_map(B,s,A,r); \
\ f O g = id(B); g O f = id(A) |] ==> f: ord_iso(A,r,B,s)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (REPEAT_FIRST (ares_tac [fg_imp_bijective]));
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (subgoal_tac "<g`(f`x), g`(f`y)> : r" 1);
-by (fast_tac (ZF_cs addIs [apply_type] addSEs [bspec RS bspec RS mp]) 2);
-by (asm_full_simp_tac
- (ZF_ss addsimps [comp_eq_id_iff RS iffD1]) 1);
+by (fast_tac (!claset addIs [apply_type] addSEs [bspec RS bspec RS mp]) 2);
+by (asm_full_simp_tac (!simpset addsimps [comp_eq_id_iff RS iffD1]) 1);
qed "mono_ord_isoI";
goal Order.thy
@@ -303,37 +299,37 @@
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, ord_iso_def]
"!!A B r. [| part_ord(B,s); f: ord_iso(A,r,B,s) |] ==> part_ord(A,r)";
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (fast_tac (ZF_cs addIs [bij_is_fun RS apply_type]) 1);
qed "part_ord_ord_iso";
goalw Order.thy [linear_def, ord_iso_def]
"!!A B r. [| linear(B,s); f: ord_iso(A,r,B,s) |] ==> linear(A,r)";
-by (asm_simp_tac ZF_ss 1);
-by (safe_tac ZF_cs);
+by (Asm_simp_tac 1);
+by (safe_tac (!claset));
by (dres_inst_tac [("x1", "f`x"), ("x", "f`xa")] (bspec RS bspec) 1);
-by (safe_tac (ZF_cs addSEs [bij_is_fun RS apply_type]));
+by (safe_tac (!claset addSEs [bij_is_fun RS apply_type]));
by (dres_inst_tac [("t", "op `(converse(f))")] subst_context 1);
-by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
+by (asm_full_simp_tac (!simpset addsimps [left_inverse_bij]) 1);
qed "linear_ord_iso";
goalw Order.thy [wf_on_def, wf_def, ord_iso_def]
"!!A B r. [| wf[B](s); f: ord_iso(A,r,B,s) |] ==> wf[A](r)";
(*reversed &-congruence rule handles context of membership in A*)
-by (asm_full_simp_tac (ZF_ss addcongs [conj_cong2]) 1);
-by (safe_tac ZF_cs);
+by (asm_full_simp_tac (!simpset addcongs [conj_cong2]) 1);
+by (safe_tac (!claset));
by (dres_inst_tac [("x", "{f`z. z:Z Int A}")] spec 1);
by (safe_tac eq_cs);
by (dtac equalityD1 1);
-by (fast_tac (ZF_cs addSIs [bexI]) 1);
-by (fast_tac (ZF_cs addSIs [bexI]
+by (fast_tac (!claset addSIs [bexI]) 1);
+by (fast_tac (!claset addSIs [bexI]
addIs [bij_is_fun RS apply_type]) 1);
qed "wf_on_ord_iso";
goalw Order.thy [well_ord_def, tot_ord_def]
"!!A B r. [| well_ord(B,s); f: ord_iso(A,r,B,s) |] ==> well_ord(A,r)";
by (fast_tac
- (ZF_cs addSEs [part_ord_ord_iso, linear_ord_iso, wf_on_ord_iso]) 1);
+ (!claset addSEs [part_ord_ord_iso, linear_ord_iso, wf_on_ord_iso]) 1);
qed "well_ord_ord_iso";
@@ -348,7 +344,7 @@
by (wf_on_ind_tac "y" [] 1);
by (dres_inst_tac [("a","y1")] (bij_is_fun RS apply_type) 1);
by (assume_tac 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "well_ord_iso_subset_lemma";
(*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment
@@ -361,7 +357,7 @@
by (EVERY1 [dtac (ord_iso_is_bij RS bij_is_fun RS apply_type),
assume_tac]);
(*Now we also know f`x : pred(A,x,r); contradiction! *)
-by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, pred_def]) 1);
+by (asm_full_simp_tac (!simpset addsimps [well_ord_def, pred_def]) 1);
qed "well_ord_iso_predE";
(*Simple consequence of Lemma 6.1*)
@@ -375,8 +371,8 @@
by (REPEAT (*because there are two symmetric cases*)
(EVERY [eresolve_tac [pred_subset RSN (2, well_ord_subset) RS
well_ord_iso_predE] 1,
- fast_tac (ZF_cs addSIs [predI]) 2,
- asm_simp_tac (ZF_ss addsimps [trans_pred_pred_eq]) 1]));
+ fast_tac (!claset addSIs [predI]) 2,
+ asm_simp_tac (!simpset addsimps [trans_pred_pred_eq]) 1]));
qed "well_ord_iso_pred_eq";
(*Does not assume r is a wellordering!*)
@@ -385,10 +381,10 @@
\ f `` pred(A,a,r) = pred(B, f`a, s)";
by (etac CollectE 1);
by (asm_simp_tac
- (ZF_ss addsimps [[bij_is_fun, Collect_subset] MRS image_fun]) 1);
+ (!simpset addsimps [[bij_is_fun, Collect_subset] MRS image_fun]) 1);
by (safe_tac (eq_cs addSEs [bij_is_fun RS apply_type]));
by (rtac RepFun_eqI 1);
-by (fast_tac (ZF_cs addSIs [right_inverse_bij RS sym]) 1);
+by (fast_tac (!claset addSIs [right_inverse_bij RS sym]) 1);
by (asm_simp_tac bij_inverse_ss 1);
qed "ord_iso_image_pred";
@@ -397,11 +393,11 @@
goal Order.thy
"!!r. [| f : ord_iso(A,r,B,s); a:A |] ==> \
\ restrict(f, pred(A,a,r)) : ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)";
-by (asm_simp_tac (ZF_ss addsimps [ord_iso_image_pred RS sym]) 1);
+by (asm_simp_tac (!simpset addsimps [ord_iso_image_pred RS sym]) 1);
by (rewtac ord_iso_def);
by (etac CollectE 1);
by (rtac CollectI 1);
-by (asm_full_simp_tac (ZF_ss addsimps [pred_def]) 2);
+by (asm_full_simp_tac (!simpset addsimps [pred_def]) 2);
by (eresolve_tac [[bij_is_inj, pred_subset] MRS restrict_bij] 1);
qed "ord_iso_restrict_pred";
@@ -414,17 +410,20 @@
by (forward_tac [ord_iso_is_bij RS bij_is_fun RS apply_type] 1 THEN
REPEAT1 (eresolve_tac [asm_rl, predI, predE] 1));
by (subgoal_tac "b = g`a" 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (rtac well_ord_iso_pred_eq 1);
by (REPEAT_SOME assume_tac);
by (forward_tac [ord_iso_restrict_pred] 1 THEN
REPEAT1 (eresolve_tac [asm_rl, predI] 1));
by (asm_full_simp_tac
- (ZF_ss addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1);
+ (!simpset addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1);
by (eresolve_tac [ord_iso_sym RS ord_iso_trans] 1);
by (assume_tac 1);
qed "well_ord_iso_preserving";
+val bij_apply_cs = !claset addSEs [bij_converse_bij, ord_iso_is_bij]
+ addIs [bij_is_fun, apply_type];
+
(*See Halmos, page 72*)
goal Order.thy
"!!r. [| well_ord(A,r); \
@@ -433,9 +432,9 @@
by (forward_tac [well_ord_iso_subset_lemma] 1);
by (res_inst_tac [("f","converse(f)"), ("g","g")] ord_iso_trans 1);
by (REPEAT_FIRST (ares_tac [subset_refl, ord_iso_sym]));
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (forward_tac [ord_iso_converse] 1);
-by (REPEAT (fast_tac (bij_apply_cs addSEs [ord_iso_is_bij]) 1));
+by (REPEAT (fast_tac bij_apply_cs 1));
by (asm_full_simp_tac bij_inverse_ss 1);
qed "well_ord_iso_unique_lemma";
@@ -446,9 +445,9 @@
by (rtac fun_extension 1);
by (REPEAT (etac (ord_iso_is_bij RS bij_is_fun) 1));
by (subgoals_tac ["f`x : B", "g`x : B", "linear(B,s)"] 1);
-by (REPEAT (fast_tac (bij_apply_cs addSEs [ord_iso_is_bij]) 3));
+by (REPEAT (fast_tac bij_apply_cs 3));
by (dtac well_ord_ord_iso 2 THEN etac ord_iso_sym 2);
-by (asm_full_simp_tac (ZF_ss addsimps [tot_ord_def, well_ord_def]) 2);
+by (asm_full_simp_tac (!simpset addsimps [tot_ord_def, well_ord_def]) 2);
by (linear_case_tac 1);
by (DEPTH_SOLVE (eresolve_tac [asm_rl, well_ord_iso_unique_lemma RS notE] 1));
qed "well_ord_iso_unique";
@@ -458,17 +457,17 @@
goalw Order.thy [ord_iso_map_def]
"ord_iso_map(A,r,B,s) <= A*B";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "ord_iso_map_subset";
goalw Order.thy [ord_iso_map_def]
"domain(ord_iso_map(A,r,B,s)) <= A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "domain_ord_iso_map";
goalw Order.thy [ord_iso_map_def]
"range(ord_iso_map(A,r,B,s)) <= B";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "range_ord_iso_map";
goalw Order.thy [ord_iso_map_def]
@@ -478,7 +477,7 @@
goalw Order.thy [ord_iso_map_def, function_def]
"!!B. well_ord(B,s) ==> function(ord_iso_map(A,r,B,s))";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac well_ord_iso_pred_eq 1);
by (REPEAT_SOME assume_tac);
by (eresolve_tac [ord_iso_sym RS ord_iso_trans] 1);
@@ -489,7 +488,7 @@
"!!B. well_ord(B,s) ==> ord_iso_map(A,r,B,s) \
\ : domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))";
by (asm_simp_tac
- (ZF_ss addsimps [Pi_iff, function_ord_iso_map,
+ (!simpset addsimps [Pi_iff, function_ord_iso_map,
ord_iso_map_subset RS domain_times_range]) 1);
qed "ord_iso_map_fun";
@@ -497,14 +496,14 @@
"!!B. [| well_ord(A,r); well_ord(B,s) |] ==> ord_iso_map(A,r,B,s) \
\ : mono_map(domain(ord_iso_map(A,r,B,s)), r, \
\ range(ord_iso_map(A,r,B,s)), s)";
-by (asm_simp_tac (ZF_ss addsimps [ord_iso_map_fun]) 1);
-by (safe_tac ZF_cs);
+by (asm_simp_tac (!simpset addsimps [ord_iso_map_fun]) 1);
+by (safe_tac (!claset));
by (subgoals_tac ["x:A", "xa:A", "y:B", "ya:B"] 1);
by (REPEAT
- (fast_tac (ZF_cs addSEs [ord_iso_map_subset RS subsetD RS SigmaE]) 2));
-by (asm_simp_tac (ZF_ss addsimps [ord_iso_map_fun RSN (2,apply_equality)]) 1);
+ (fast_tac (!claset addSEs [ord_iso_map_subset RS subsetD RS SigmaE]) 2));
+by (asm_simp_tac (!simpset addsimps [ord_iso_map_fun RSN (2,apply_equality)]) 1);
by (rewtac ord_iso_map_def);
-by (safe_tac (ZF_cs addSEs [UN_E]));
+by (safe_tac (!claset addSEs [UN_E]));
by (rtac well_ord_iso_preserving 1 THEN REPEAT_FIRST assume_tac);
qed "ord_iso_map_mono_map";
@@ -515,7 +514,7 @@
by (rtac well_ord_mono_ord_isoI 1);
by (resolve_tac [converse_ord_iso_map RS subst] 4);
by (asm_simp_tac
- (ZF_ss addsimps [ord_iso_map_subset RS converse_converse]) 4);
+ (!simpset addsimps [ord_iso_map_subset RS converse_converse]) 4);
by (REPEAT (ares_tac [ord_iso_map_mono_map] 3));
by (ALLGOALS (etac well_ord_subset));
by (ALLGOALS (resolve_tac [domain_ord_iso_map, range_ord_iso_map]));
@@ -526,20 +525,20 @@
"!!B. [| well_ord(A,r); well_ord(B,s); \
\ a: A; a ~: domain(ord_iso_map(A,r,B,s)) \
\ |] ==> domain(ord_iso_map(A,r,B,s)) <= pred(A, a, r)";
-by (safe_tac (ZF_cs addSIs [predI]));
+by (safe_tac (!claset addSIs [predI]));
(*Case analysis on xaa vs a in r *)
by (forw_inst_tac [("A","A")] well_ord_is_linear 1);
by (linear_case_tac 1);
(*Trivial case: a=xa*)
-by (fast_tac ZF_cs 2);
+by (Fast_tac 2);
(*Harder case: <a, xa>: r*)
by (forward_tac [ord_iso_is_bij RS bij_is_fun RS apply_type] 1 THEN
REPEAT1 (eresolve_tac [asm_rl, predI, predE] 1));
by (forward_tac [ord_iso_restrict_pred] 1 THEN
REPEAT1 (eresolve_tac [asm_rl, predI] 1));
by (asm_full_simp_tac
- (ZF_ss addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1);
-by (fast_tac ZF_cs 1);
+ (!simpset addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1);
+by (Fast_tac 1);
qed "domain_ord_iso_map_subset";
(*For the 4-way case analysis in the main result*)
@@ -550,7 +549,7 @@
by (forward_tac [well_ord_is_wf] 1);
by (rewrite_goals_tac [wf_on_def, wf_def]);
by (dres_inst_tac [("x", "A-domain(ord_iso_map(A,r,B,s))")] spec 1);
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
(*The first case: the domain equals A*)
by (rtac (domain_ord_iso_map RS equalityI) 1);
by (etac (Diff_eq_0_iff RS iffD1) 1);
@@ -558,8 +557,8 @@
by (swap_res_tac [bexI] 1);
by (assume_tac 2);
by (rtac equalityI 1);
-(*not ZF_cs below; that would use rules like domainE!*)
-by (fast_tac (pair_cs addSEs [predE]) 2);
+(*not (!claset) below; that would use rules like domainE!*)
+by (fast_tac (!claset addSEs [predE]) 2);
by (REPEAT (ares_tac [domain_ord_iso_map_subset] 1));
qed "domain_ord_iso_map_cases";
@@ -570,7 +569,7 @@
\ (EX y:B. range(ord_iso_map(A,r,B,s))= pred(B,y,s))";
by (resolve_tac [converse_ord_iso_map RS subst] 1);
by (asm_simp_tac
- (ZF_ss addsimps [range_converse, domain_ord_iso_map_cases]) 1);
+ (!simpset addsimps [range_converse, domain_ord_iso_map_cases]) 1);
qed "range_ord_iso_map_cases";
(*Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets*)
@@ -583,15 +582,15 @@
by (forw_inst_tac [("B","B")] range_ord_iso_map_cases 2);
by (REPEAT_FIRST (eresolve_tac [asm_rl, disjE, bexE]));
by (ALLGOALS (dtac ord_iso_map_ord_iso THEN' assume_tac THEN'
- asm_full_simp_tac (ZF_ss addsimps [bexI])));
+ asm_full_simp_tac (!simpset addsimps [bexI])));
by (resolve_tac [wf_on_not_refl RS notE] 1);
by (etac well_ord_is_wf 1);
by (assume_tac 1);
by (subgoal_tac "<x,y>: ord_iso_map(A,r,B,s)" 1);
by (dtac rangeI 1);
-by (asm_full_simp_tac (ZF_ss addsimps [pred_def]) 1);
+by (asm_full_simp_tac (!simpset addsimps [pred_def]) 1);
by (rewtac ord_iso_map_def);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "well_ord_trichotomy";
@@ -599,25 +598,50 @@
goalw Order.thy [irrefl_def]
"!!A. irrefl(A,r) ==> irrefl(A,converse(r))";
-by (fast_tac (ZF_cs addSIs [converseI]) 1);
+by (fast_tac (!claset addSIs [converseI]) 1);
qed "irrefl_converse";
goalw Order.thy [trans_on_def]
"!!A. trans[A](r) ==> trans[A](converse(r))";
-by (fast_tac (ZF_cs addSIs [converseI]) 1);
+by (fast_tac (!claset addSIs [converseI]) 1);
qed "trans_on_converse";
goalw Order.thy [part_ord_def]
"!!A. part_ord(A,r) ==> part_ord(A,converse(r))";
-by (fast_tac (ZF_cs addSIs [irrefl_converse, trans_on_converse]) 1);
+by (fast_tac (!claset addSIs [irrefl_converse, trans_on_converse]) 1);
qed "part_ord_converse";
goalw Order.thy [linear_def]
"!!A. linear(A,r) ==> linear(A,converse(r))";
-by (fast_tac (ZF_cs addSIs [converseI]) 1);
+by (fast_tac (!claset addSIs [converseI]) 1);
qed "linear_converse";
goalw Order.thy [tot_ord_def]
"!!A. tot_ord(A,r) ==> tot_ord(A,converse(r))";
-by (fast_tac (ZF_cs addSIs [part_ord_converse, linear_converse]) 1);
+by (fast_tac (!claset addSIs [part_ord_converse, linear_converse]) 1);
qed "tot_ord_converse";
+
+
+(** By Krzysztof Grabczewski.
+ Lemmas involving the first element of a well ordered set **)
+
+goalw thy [first_def] "!!b. first(b,B,r) ==> b:B";
+by (Fast_tac 1);
+qed "first_is_elem";
+
+goalw thy [well_ord_def, wf_on_def, wf_def, first_def]
+ "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (EX! b. first(b,B,r))";
+by (REPEAT (eresolve_tac [conjE,allE,disjE] 1));
+by (contr_tac 1);
+by (etac bexE 1);
+by (res_inst_tac [("a","x")] ex1I 1);
+by (Fast_tac 2);
+by (rewrite_goals_tac [tot_ord_def, linear_def]);
+by (Fast_tac 1);
+qed "well_ord_imp_ex1_first";
+
+goal thy "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B";
+by (dtac well_ord_imp_ex1_first 1 THEN REPEAT (assume_tac 1));
+by (rtac first_is_elem 1);
+by (etac theI 1);
+qed "the_first_in";
--- a/src/ZF/Order.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Order.thy Fri Jan 03 15:01:55 1997 +0100
@@ -38,4 +38,9 @@
UN x:A. UN y:B. UN f: ord_iso(pred(A,x,r), r, pred(B,y,s), s).
{<x,y>}"
+constdefs
+
+ first :: [i, i, i] => o
+ "first(u, X, R) == u:X & (ALL v:X. v~=u --> <u,v> : R)"
+
end
--- a/src/ZF/OrderArith.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/OrderArith.ML Fri Jan 03 15:01:55 1997 +0100
@@ -58,17 +58,17 @@
by (rtac Collect_subset 1);
qed "radd_type";
-bind_thm ("field_radd", (radd_type RS field_rel_subset));
+bind_thm ("field_radd", radd_type RS field_rel_subset);
(** Linearity **)
-val radd_ss = sum_ss addsimps [radd_Inl_iff, radd_Inr_iff,
- radd_Inl_Inr_iff, radd_Inr_Inl_iff];
+Addsimps [radd_Inl_iff, radd_Inr_iff,
+ radd_Inl_Inr_iff, radd_Inr_Inl_iff];
goalw OrderArith.thy [linear_def]
"!!r s. [| linear(A,r); linear(B,s) |] ==> linear(A+B,radd(A,r,B,s))";
by (REPEAT_FIRST (ares_tac [ballI] ORELSE' etac sumE));
-by (ALLGOALS (asm_simp_tac radd_ss));
+by (ALLGOALS Asm_simp_tac);
qed "linear_radd";
@@ -96,7 +96,7 @@
goal OrderArith.thy
"!!r s. [| wf(r); wf(s) |] ==> wf(radd(field(r),r,field(s),s))";
-by (asm_full_simp_tac (ZF_ss addsimps [wf_iff_wf_on_field]) 1);
+by (asm_full_simp_tac (!simpset addsimps [wf_iff_wf_on_field]) 1);
by (rtac (field_radd RSN (2, wf_on_subset_A)) 1);
by (REPEAT (ares_tac [wf_on_radd] 1));
qed "wf_radd";
@@ -105,39 +105,35 @@
"!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \
\ well_ord(A+B, radd(A,r,B,s))";
by (rtac well_ordI 1);
-by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, wf_on_radd]) 1);
+by (asm_full_simp_tac (!simpset addsimps [well_ord_def, wf_on_radd]) 1);
by (asm_full_simp_tac
- (ZF_ss addsimps [well_ord_def, tot_ord_def, linear_radd]) 1);
+ (!simpset addsimps [well_ord_def, tot_ord_def, linear_radd]) 1);
qed "well_ord_radd";
(** An ord_iso congruence law **)
-val case_ss =
- bij_inverse_ss addsimps [Inl_iff, Inl_Inr_iff, Inr_iff, Inr_Inl_iff,
- case_Inl, case_Inr, InlI, InrI];
-
goal OrderArith.thy
"!!f g. [| f: bij(A,C); g: bij(B,D) |] ==> \
\ (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) : bij(A+B, C+D)";
by (res_inst_tac
[("d", "case(%x. Inl(converse(f)`x), %y. Inr(converse(g)`y))")]
lam_bijective 1);
-by (safe_tac (ZF_cs addSEs [sumE]));
-by (ALLGOALS (asm_simp_tac case_ss));
+by (safe_tac (!claset addSEs [sumE]));
+by (ALLGOALS (asm_simp_tac bij_inverse_ss));
qed "sum_bij";
goalw OrderArith.thy [ord_iso_def]
"!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \
\ (lam z:A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \
\ : ord_iso(A+B, radd(A,r,B,s), A'+B', radd(A',r',B',s'))";
-by (safe_tac (ZF_cs addSIs [sum_bij]));
+by (safe_tac (!claset addSIs [sum_bij]));
(*Do the beta-reductions now*)
-by (ALLGOALS (asm_full_simp_tac ZF_ss));
+by (ALLGOALS (Asm_full_simp_tac));
by (safe_tac sum_cs);
(*8 subgoals!*)
by (ALLGOALS
(asm_full_simp_tac
- (radd_ss addcongs [conj_cong] addsimps [bij_is_fun RS apply_type])));
+ (!simpset addcongs [conj_cong] addsimps [bij_is_fun RS apply_type])));
qed "sum_ord_iso_cong";
(*Could we prove an ord_iso result? Perhaps
@@ -150,8 +146,8 @@
by (fast_tac (sum_cs addSIs [if_type]) 2);
by (DEPTH_SOLVE_1 (eresolve_tac [case_type, UnI1, UnI2] 1));
by (safe_tac sum_cs);
-by (ALLGOALS (asm_simp_tac (sum_ss setloop split_tac [expand_if])));
-by (fast_tac (ZF_cs addEs [equalityE]) 1);
+by (ALLGOALS (asm_simp_tac (!simpset setloop split_tac [expand_if])));
+by (fast_tac (!claset addEs [equalityE]) 1);
qed "sum_disjoint_bij";
(** Associativity **)
@@ -161,7 +157,7 @@
\ : bij((A+B)+C, A+(B+C))";
by (res_inst_tac [("d", "case(%x.Inl(Inl(x)), case(%x.Inl(Inr(x)), Inr))")]
lam_bijective 1);
-by (ALLGOALS (asm_simp_tac (case_ss setloop etac sumE)));
+by (ALLGOALS (asm_simp_tac (!simpset setloop etac sumE)));
qed "sum_assoc_bij";
goal OrderArith.thy
@@ -170,7 +166,7 @@
\ A+(B+C), radd(A, r, B+C, radd(B,s,C,t)))";
by (resolve_tac [sum_assoc_bij RS ord_isoI] 1);
by (REPEAT_FIRST (etac sumE));
-by (ALLGOALS (asm_simp_tac radd_ss));
+by (ALLGOALS Asm_simp_tac);
qed "sum_assoc_ord_iso";
@@ -182,9 +178,11 @@
"!!r s. <<a',b'>, <a,b>> : rmult(A,r,B,s) <-> \
\ (<a',a>: r & a':A & a:A & b': B & b: B) | \
\ (<b',b>: s & a'=a & a:A & b': B & b: B)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "rmult_iff";
+Addsimps [rmult_iff];
+
val major::prems = goal OrderArith.thy
"[| <<a',b'>, <a,b>> : rmult(A,r,B,s); \
\ [| <a',a>: r; a':A; a:A; b':B; b:B |] ==> Q; \
@@ -208,10 +206,10 @@
"[| linear(A,r); linear(B,s) |] ==> linear(A*B,rmult(A,r,B,s))";
by (rewtac linear_def); (*Note! the premises are NOT rewritten*)
by (REPEAT_FIRST (ares_tac [ballI] ORELSE' etac SigmaE));
-by (asm_simp_tac (ZF_ss addsimps [rmult_iff]) 1);
+by (Asm_simp_tac 1);
by (res_inst_tac [("x","xa"), ("y","xb")] (lina RS linearE) 1);
by (res_inst_tac [("x","ya"), ("y","yb")] (linb RS linearE) 4);
-by (REPEAT_SOME (fast_tac ZF_cs));
+by (REPEAT_SOME (Fast_tac));
qed "linear_rmult";
@@ -223,19 +221,19 @@
by (etac SigmaE 1);
by (etac ssubst 1);
by (subgoal_tac "ALL b:B. <x,b>: Ba" 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (eres_inst_tac [("a","x")] wf_on_induct 1 THEN assume_tac 1);
by (rtac ballI 1);
by (eres_inst_tac [("a","b")] wf_on_induct 1 THEN assume_tac 1);
by (etac (bspec RS mp) 1);
-by (fast_tac ZF_cs 1);
-by (best_tac (ZF_cs addSEs [rmultE]) 1);
+by (Fast_tac 1);
+by (best_tac (!claset addSEs [rmultE]) 1);
qed "wf_on_rmult";
goal OrderArith.thy
"!!r s. [| wf(r); wf(s) |] ==> wf(rmult(field(r),r,field(s),s))";
-by (asm_full_simp_tac (ZF_ss addsimps [wf_iff_wf_on_field]) 1);
+by (asm_full_simp_tac (!simpset addsimps [wf_iff_wf_on_field]) 1);
by (rtac (field_rmult RSN (2, wf_on_subset_A)) 1);
by (REPEAT (ares_tac [wf_on_rmult] 1));
qed "wf_rmult";
@@ -244,9 +242,9 @@
"!!r s. [| well_ord(A,r); well_ord(B,s) |] ==> \
\ well_ord(A*B, rmult(A,r,B,s))";
by (rtac well_ordI 1);
-by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, wf_on_rmult]) 1);
+by (asm_full_simp_tac (!simpset addsimps [well_ord_def, wf_on_rmult]) 1);
by (asm_full_simp_tac
- (ZF_ss addsimps [well_ord_def, tot_ord_def, linear_rmult]) 1);
+ (!simpset addsimps [well_ord_def, tot_ord_def, linear_rmult]) 1);
qed "well_ord_rmult";
@@ -257,7 +255,7 @@
\ (lam <x,y>:A*B. <f`x, g`y>) : bij(A*B, C*D)";
by (res_inst_tac [("d", "%<x,y>. <converse(f)`x, converse(g)`y>")]
lam_bijective 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (ALLGOALS (asm_simp_tac bij_inverse_ss));
qed "prod_bij";
@@ -265,18 +263,17 @@
"!!r s. [| f: ord_iso(A,r,A',r'); g: ord_iso(B,s,B',s') |] ==> \
\ (lam <x,y>:A*B. <f`x, g`y>) \
\ : ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))";
-by (safe_tac (ZF_cs addSIs [prod_bij]));
+by (safe_tac (!claset addSIs [prod_bij]));
by (ALLGOALS
- (asm_full_simp_tac
- (ZF_ss addsimps [rmult_iff, bij_is_fun RS apply_type])));
-by (fast_tac ZF_cs 1);
-by (fast_tac (ZF_cs addSEs [bij_is_inj RS inj_apply_equality]) 1);
+ (asm_full_simp_tac (!simpset addsimps [bij_is_fun RS apply_type])));
+by (Fast_tac 1);
+by (fast_tac (!claset addSEs [bij_is_inj RS inj_apply_equality]) 1);
qed "prod_ord_iso_cong";
goal OrderArith.thy "(lam z:A. <x,z>) : bij(A, {x}*A)";
by (res_inst_tac [("d", "snd")] lam_bijective 1);
-by (safe_tac ZF_cs);
-by (ALLGOALS (asm_simp_tac ZF_ss));
+by (safe_tac (!claset));
+by (ALLGOALS Asm_simp_tac);
qed "singleton_prod_bij";
(*Used??*)
@@ -284,8 +281,8 @@
"!!x xr. well_ord({x},xr) ==> \
\ (lam z:A. <x,z>) : ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))";
by (resolve_tac [singleton_prod_bij RS ord_isoI] 1);
-by (asm_simp_tac (ZF_ss addsimps [rmult_iff]) 1);
-by (fast_tac (ZF_cs addEs [well_ord_is_wf RS wf_on_not_refl RS notE]) 1);
+by (Asm_simp_tac 1);
+by (fast_tac (!claset addEs [well_ord_is_wf RS wf_on_not_refl RS notE]) 1);
qed "singleton_prod_ord_iso";
(*Here we build a complicated function term, then simplify it using
@@ -299,10 +296,10 @@
by (rtac singleton_prod_bij 1);
by (rtac sum_disjoint_bij 1);
by (fast_tac eq_cs 1);
-by (asm_simp_tac (ZF_ss addcongs [case_cong] addsimps [id_conv]) 1);
+by (asm_simp_tac (!simpset addcongs [case_cong] addsimps [id_conv]) 1);
by (resolve_tac [comp_lam RS trans RS sym] 1);
by (fast_tac (sum_cs addSEs [case_type]) 1);
-by (asm_simp_tac (ZF_ss addsimps [case_case]) 1);
+by (asm_simp_tac (!simpset addsimps [case_case]) 1);
qed "prod_sum_singleton_bij";
goal OrderArith.thy
@@ -313,11 +310,11 @@
\ pred(A,a,r)*B Un {a}*pred(B,b,s), rmult(A,r,B,s))";
by (resolve_tac [prod_sum_singleton_bij RS ord_isoI] 1);
by (asm_simp_tac
- (ZF_ss addsimps [pred_iff, well_ord_is_wf RS wf_on_not_refl]) 1);
-by (asm_simp_tac ZF_ss 1);
+ (!simpset addsimps [pred_iff, well_ord_is_wf RS wf_on_not_refl]) 1);
+by (Asm_simp_tac 1);
by (REPEAT_FIRST (eresolve_tac [SigmaE, sumE, predE]));
-by (ALLGOALS (asm_simp_tac (radd_ss addsimps [rmult_iff])));
-by (ALLGOALS (fast_tac (ZF_cs addEs [well_ord_is_wf RS wf_on_asym])));
+by (ALLGOALS (Asm_simp_tac));
+by (ALLGOALS (fast_tac (!claset addEs [well_ord_is_wf RS wf_on_asym])));
qed "prod_sum_singleton_ord_iso";
(** Distributive law **)
@@ -327,8 +324,8 @@
\ : bij((A+B)*C, (A*C)+(B*C))";
by (res_inst_tac
[("d", "case(%<x,y>.<Inl(x),y>, %<x,y>.<Inr(x),y>)")] lam_bijective 1);
-by (safe_tac (ZF_cs addSEs [sumE]));
-by (ALLGOALS (asm_simp_tac case_ss));
+by (safe_tac (!claset addSEs [sumE]));
+by (ALLGOALS Asm_simp_tac);
qed "sum_prod_distrib_bij";
goal OrderArith.thy
@@ -337,7 +334,7 @@
\ (A*C)+(B*C), radd(A*C, rmult(A,r,C,t), B*C, rmult(B,s,C,t)))";
by (resolve_tac [sum_prod_distrib_bij RS ord_isoI] 1);
by (REPEAT_FIRST (eresolve_tac [SigmaE, sumE]));
-by (ALLGOALS (asm_simp_tac (radd_ss addsimps [rmult_iff])));
+by (ALLGOALS (Asm_simp_tac));
qed "sum_prod_distrib_ord_iso";
(** Associativity **)
@@ -345,7 +342,7 @@
goal OrderArith.thy
"(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>) : bij((A*B)*C, A*(B*C))";
by (res_inst_tac [("d", "%<x, <y,z>>. <<x,y>, z>")] lam_bijective 1);
-by (ALLGOALS (asm_simp_tac (case_ss setloop etac SigmaE)));
+by (ALLGOALS (asm_simp_tac (!simpset setloop etac SigmaE)));
qed "prod_assoc_bij";
goal OrderArith.thy
@@ -354,8 +351,8 @@
\ A*(B*C), rmult(A, r, B*C, rmult(B,s,C,t)))";
by (resolve_tac [prod_assoc_bij RS ord_isoI] 1);
by (REPEAT_FIRST (eresolve_tac [SigmaE, ssubst]));
-by (asm_simp_tac (ZF_ss addsimps [rmult_iff]) 1);
-by (fast_tac ZF_cs 1);
+by (Asm_simp_tac 1);
+by (Fast_tac 1);
qed "prod_assoc_ord_iso";
(**** Inverse image of a relation ****)
@@ -364,7 +361,7 @@
goalw OrderArith.thy [rvimage_def]
"<a,b> : rvimage(A,f,r) <-> <f`a,f`b>: r & a:A & b:A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "rvimage_iff";
(** Type checking **)
@@ -385,17 +382,17 @@
goalw OrderArith.thy [irrefl_def, rvimage_def]
"!!A B. [| f: inj(A,B); irrefl(B,r) |] ==> irrefl(A, rvimage(A,f,r))";
-by (fast_tac (ZF_cs addIs [inj_is_fun RS apply_type]) 1);
+by (fast_tac (!claset addIs [inj_is_fun RS apply_type]) 1);
qed "irrefl_rvimage";
goalw OrderArith.thy [trans_on_def, rvimage_def]
"!!A B. [| f: inj(A,B); trans[B](r) |] ==> trans[A](rvimage(A,f,r))";
-by (fast_tac (ZF_cs addIs [inj_is_fun RS apply_type]) 1);
+by (fast_tac (!claset addIs [inj_is_fun RS apply_type]) 1);
qed "trans_on_rvimage";
goalw OrderArith.thy [part_ord_def]
"!!A B. [| f: inj(A,B); part_ord(B,r) |] ==> part_ord(A, rvimage(A,f,r))";
-by (fast_tac (ZF_cs addSIs [irrefl_rvimage, trans_on_rvimage]) 1);
+by (fast_tac (!claset addSIs [irrefl_rvimage, trans_on_rvimage]) 1);
qed "part_ord_rvimage";
(** Linearity **)
@@ -404,15 +401,15 @@
"[| f: inj(A,B); linear(B,r) |] ==> linear(A,rvimage(A,f,r))";
by (rewtac linear_def); (*Note! the premises are NOT rewritten*)
by (REPEAT_FIRST (ares_tac [ballI]));
-by (asm_simp_tac (ZF_ss addsimps [rvimage_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [rvimage_iff]) 1);
by (cut_facts_tac [finj] 1);
by (res_inst_tac [("x","f`x"), ("y","f`y")] (lin RS linearE) 1);
-by (REPEAT_SOME (fast_tac (ZF_cs addSIs [apply_type])));
+by (REPEAT_SOME (fast_tac (!claset addSIs [apply_type])));
qed "linear_rvimage";
goalw OrderArith.thy [tot_ord_def]
"!!A B. [| f: inj(A,B); tot_ord(B,r) |] ==> tot_ord(A, rvimage(A,f,r))";
-by (fast_tac (ZF_cs addSIs [part_ord_rvimage, linear_rvimage]) 1);
+by (fast_tac (!claset addSIs [part_ord_rvimage, linear_rvimage]) 1);
qed "tot_ord_rvimage";
@@ -422,10 +419,10 @@
"!!r. [| f: A->B; wf[B](r) |] ==> wf[A](rvimage(A,f,r))";
by (rtac wf_onI2 1);
by (subgoal_tac "ALL z:A. f`z=f`y --> z: Ba" 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (eres_inst_tac [("a","f`y")] wf_on_induct 1);
-by (fast_tac (ZF_cs addSIs [apply_type]) 1);
-by (best_tac (ZF_cs addSIs [apply_type] addSDs [rvimage_iff RS iffD1]) 1);
+by (fast_tac (!claset addSIs [apply_type]) 1);
+by (best_tac (!claset addSIs [apply_type] addSDs [rvimage_iff RS iffD1]) 1);
qed "wf_on_rvimage";
(*Note that we need only wf[A](...) and linear(A,...) to get the result!*)
@@ -433,13 +430,13 @@
"!!r. [| f: inj(A,B); well_ord(B,r) |] ==> well_ord(A, rvimage(A,f,r))";
by (rtac well_ordI 1);
by (rewrite_goals_tac [well_ord_def, tot_ord_def]);
-by (fast_tac (ZF_cs addSIs [wf_on_rvimage, inj_is_fun]) 1);
-by (fast_tac (ZF_cs addSIs [linear_rvimage]) 1);
+by (fast_tac (!claset addSIs [wf_on_rvimage, inj_is_fun]) 1);
+by (fast_tac (!claset addSIs [linear_rvimage]) 1);
qed "well_ord_rvimage";
goalw OrderArith.thy [ord_iso_def]
"!!A B. f: bij(A,B) ==> f: ord_iso(A, rvimage(A,f,s), B, s)";
-by (asm_full_simp_tac (ZF_ss addsimps [rvimage_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [rvimage_iff]) 1);
qed "ord_iso_rvimage";
goalw OrderArith.thy [ord_iso_def, rvimage_def]
--- a/src/ZF/OrderType.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/OrderType.ML Fri Jan 03 15:01:55 1997 +0100
@@ -18,7 +18,7 @@
by (rtac well_ordI 1);
by (rtac (wf_Memrel RS wf_imp_wf_on) 1);
by (resolve_tac [prem RS ltE] 1);
-by (asm_simp_tac (ZF_ss addsimps [linear_def, Memrel_iff,
+by (asm_simp_tac (!simpset addsimps [linear_def, Memrel_iff,
[ltI, prem] MRS lt_trans2 RS ltD]) 1);
by (REPEAT (resolve_tac [ballI, Ord_linear] 1));
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1));
@@ -31,7 +31,7 @@
The smaller ordinal is an initial segment of the larger *)
goalw OrderType.thy [pred_def, lt_def]
"!!i j. j<i ==> pred(i, j, Memrel(i)) = j";
-by (asm_simp_tac (ZF_ss addsimps [Memrel_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [Memrel_iff]) 1);
by (fast_tac (eq_cs addEs [Ord_trans]) 1);
qed "lt_pred_Memrel";
@@ -46,10 +46,10 @@
by (etac ltE 1);
by (rtac (well_ord_Memrel RS well_ord_iso_predE) 1 THEN
assume_tac 3 THEN assume_tac 1);
-by (asm_full_simp_tac (ZF_ss addsimps [ord_iso_def]) 1);
+by (asm_full_simp_tac (!simpset addsimps [ord_iso_def]) 1);
(*Combining the two simplifications causes looping*)
-by (asm_simp_tac (ZF_ss addsimps [Memrel_iff]) 1);
-by (fast_tac (ZF_cs addSEs [bij_is_fun RS apply_type] addEs [Ord_trans]) 1);
+by (asm_simp_tac (!simpset addsimps [Memrel_iff]) 1);
+by (fast_tac (!claset addSEs [bij_is_fun RS apply_type] addEs [Ord_trans]) 1);
qed "Ord_iso_implies_eq_lemma";
(*Kunen's Theorem 7.3 (ii), page 16. Isomorphic ordinals are equal*)
@@ -76,10 +76,10 @@
goalw OrderType.thy [ordermap_def, pred_def]
"!!r. [| wf[A](r); x:A |] ==> \
\ ordermap(A,r) ` x = ordermap(A,r) `` pred(A,x,r)";
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (etac (wfrec_on RS trans) 1);
by (assume_tac 1);
-by (asm_simp_tac (ZF_ss addsimps [subset_iff, image_lam,
+by (asm_simp_tac (!simpset addsimps [subset_iff, image_lam,
vimage_singleton_iff]) 1);
qed "ordermap_eq_image";
@@ -87,7 +87,7 @@
goal OrderType.thy
"!!r. [| wf[A](r); x:A |] ==> \
\ ordermap(A,r) ` x = {ordermap(A,r)`y . y : pred(A,x,r)}";
-by (asm_simp_tac (ZF_ss addsimps [ordermap_eq_image, pred_subset,
+by (asm_simp_tac (!simpset addsimps [ordermap_eq_image, pred_subset,
ordermap_type RS image_fun]) 1);
qed "ordermap_pred_unfold";
@@ -103,26 +103,26 @@
goalw OrderType.thy [well_ord_def, tot_ord_def, part_ord_def]
"!!r. [| well_ord(A,r); x:A |] ==> Ord(ordermap(A,r) ` x)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (wf_on_ind_tac "x" [] 1);
-by (asm_simp_tac (ZF_ss addsimps [ordermap_pred_unfold]) 1);
+by (asm_simp_tac (!simpset addsimps [ordermap_pred_unfold]) 1);
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1);
by (rewrite_goals_tac [pred_def,Transset_def]);
-by (fast_tac ZF_cs 2);
-by (safe_tac ZF_cs);
+by (Fast_tac 2);
+by (safe_tac (!claset));
by (ordermap_elim_tac 1);
-by (fast_tac (ZF_cs addSEs [trans_onD]) 1);
+by (fast_tac (!claset addSEs [trans_onD]) 1);
qed "Ord_ordermap";
goalw OrderType.thy [ordertype_def]
"!!r. well_ord(A,r) ==> Ord(ordertype(A,r))";
by (stac ([ordermap_type, subset_refl] MRS image_fun) 1);
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1);
-by (fast_tac (ZF_cs addIs [Ord_ordermap]) 2);
+by (fast_tac (!claset addIs [Ord_ordermap]) 2);
by (rewrite_goals_tac [Transset_def,well_ord_def]);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (ordermap_elim_tac 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Ord_ordertype";
(*** ordermap preserves the orderings in both directions ***)
@@ -132,16 +132,16 @@
\ ordermap(A,r)`w : ordermap(A,r)`x";
by (eres_inst_tac [("x1", "x")] (ordermap_unfold RS ssubst) 1);
by (assume_tac 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "ordermap_mono";
(*linearity of r is crucial here*)
goalw OrderType.thy [well_ord_def, tot_ord_def]
"!!r. [| ordermap(A,r)`w : ordermap(A,r)`x; well_ord(A,r); \
\ w: A; x: A |] ==> <w,x>: r";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (linear_case_tac 1);
-by (fast_tac (ZF_cs addSEs [mem_not_refl RS notE]) 1);
+by (fast_tac (!claset addSEs [mem_not_refl RS notE]) 1);
by (dtac ordermap_mono 1);
by (REPEAT_SOME assume_tac);
by (etac mem_asym 1);
@@ -154,10 +154,10 @@
goalw OrderType.thy [well_ord_def, tot_ord_def, bij_def, inj_def]
"!!r. well_ord(A,r) ==> ordermap(A,r) : bij(A, ordertype(A,r))";
-by (fast_tac (ZF_cs addSIs [ordermap_type, ordermap_surj]
+by (fast_tac (!claset addSIs [ordermap_type, ordermap_surj]
addEs [linearE]
addDs [ordermap_mono]
- addss (ZF_ss addsimps [mem_not_refl])) 1);
+ addss (!simpset addsimps [mem_not_refl])) 1);
qed "ordermap_bij";
(*** Isomorphisms involving ordertype ***)
@@ -165,12 +165,12 @@
goalw OrderType.thy [ord_iso_def]
"!!r. well_ord(A,r) ==> \
\ ordermap(A,r) : ord_iso(A,r, ordertype(A,r), Memrel(ordertype(A,r)))";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (rtac ordermap_bij 1);
by (assume_tac 1);
-by (fast_tac (ZF_cs addSEs [MemrelE, converse_ordermap_mono]) 2);
+by (fast_tac (!claset addSEs [MemrelE, converse_ordermap_mono]) 2);
by (rewtac well_ord_def);
-by (fast_tac (ZF_cs addSIs [MemrelI, ordermap_mono,
+by (fast_tac (!claset addSIs [MemrelI, ordermap_mono,
ordermap_type RS apply_type]) 1);
qed "ordertype_ord_iso";
@@ -180,7 +180,7 @@
by (forward_tac [well_ord_ord_iso] 1 THEN assume_tac 1);
by (rtac Ord_iso_implies_eq 1
THEN REPEAT (etac Ord_ordertype 1));
-by (deepen_tac (ZF_cs addIs [ord_iso_trans, ord_iso_sym]
+by (deepen_tac (!claset addIs [ord_iso_trans, ord_iso_sym]
addSEs [ordertype_ord_iso]) 0 1);
qed "ordertype_eq";
@@ -205,8 +205,8 @@
by (rtac ord_iso_trans 1);
by (eresolve_tac [le_well_ord_Memrel RS ordertype_ord_iso] 2);
by (resolve_tac [id_bij RS ord_isoI] 1);
-by (asm_simp_tac (ZF_ss addsimps [id_conv, Memrel_iff]) 1);
-by (fast_tac (ZF_cs addEs [ltE, Ord_in_Ord, Ord_trans]) 1);
+by (asm_simp_tac (!simpset addsimps [id_conv, Memrel_iff]) 1);
+by (fast_tac (!claset addEs [ltE, Ord_in_Ord, Ord_trans]) 1);
qed "le_ordertype_Memrel";
(*"Ord(i) ==> ordertype(i, Memrel(i)) = i"*)
@@ -219,6 +219,8 @@
by (resolve_tac [Ord_0 RS ordertype_Memrel] 1);
qed "ordertype_0";
+Addsimps [ordertype_0];
+
(*Ordertype of rvimage: [| f: bij(A,B); well_ord(B,s) |] ==>
ordertype(A, rvimage(A,f,s)) = ordertype(B,s) *)
bind_thm ("bij_ordertype_vimage", ord_iso_rvimage RS ordertype_eq);
@@ -231,12 +233,12 @@
\ ordermap(pred(A,y,r), r) ` z = ordermap(A, r) ` z";
by (forward_tac [[well_ord_is_wf, pred_subset] MRS wf_on_subset_A] 1);
by (wf_on_ind_tac "z" [] 1);
-by (safe_tac (ZF_cs addSEs [predE]));
+by (safe_tac (!claset addSEs [predE]));
by (asm_simp_tac
- (ZF_ss addsimps [ordermap_pred_unfold, well_ord_is_wf, pred_iff]) 1);
+ (!simpset addsimps [ordermap_pred_unfold, well_ord_is_wf, pred_iff]) 1);
(*combining these two simplifications LOOPS! *)
-by (asm_simp_tac (ZF_ss addsimps [pred_pred_eq]) 1);
-by (asm_full_simp_tac (ZF_ss addsimps [pred_def]) 1);
+by (asm_simp_tac (!simpset addsimps [pred_pred_eq]) 1);
+by (asm_full_simp_tac (!simpset addsimps [pred_def]) 1);
by (rtac (refl RSN (2,RepFun_cong)) 1);
by (dtac well_ord_is_trans_on 1);
by (fast_tac (eq_cs addSEs [trans_onD]) 1);
@@ -252,9 +254,9 @@
goal OrderType.thy
"!!r. [| well_ord(A,r); x:A |] ==> \
\ ordertype(pred(A,x,r),r) <= ordertype(A,r)";
-by (asm_simp_tac (ZF_ss addsimps [ordertype_unfold,
+by (asm_simp_tac (!simpset addsimps [ordertype_unfold,
pred_subset RSN (2, well_ord_subset)]) 1);
-by (fast_tac (ZF_cs addIs [ordermap_pred_eq_ordermap, RepFun_eqI]
+by (fast_tac (!claset addIs [ordermap_pred_eq_ordermap, RepFun_eqI]
addEs [predE]) 1);
qed "ordertype_pred_subset";
@@ -275,8 +277,8 @@
\ ordertype(A,r) = {ordertype(pred(A,x,r),r). x:A}";
by (safe_tac (eq_cs addSIs [ordertype_pred_lt RS ltD]));
by (fast_tac
- (ZF_cs addss
- (ZF_ss addsimps [ordertype_def,
+ (!claset addss
+ (!simpset addsimps [ordertype_def,
well_ord_is_wf RS ordermap_eq_image,
ordermap_type RS image_fun,
ordermap_pred_eq_ordermap,
@@ -299,12 +301,12 @@
goalw OrderType.thy [Ord_alt_def, Ord_def, Transset_def, well_ord_def,
tot_ord_def, part_ord_def, trans_on_def]
"!!i. Ord_alt(i) ==> Ord(i)";
-by (asm_full_simp_tac (ZF_ss addsimps [Memrel_iff, pred_Memrel]) 1);
-by (safe_tac ZF_cs);
-by (fast_tac (ZF_cs addSDs [equalityD1]) 1);
+by (asm_full_simp_tac (!simpset addsimps [Memrel_iff, pred_Memrel]) 1);
+by (safe_tac (!claset));
+by (fast_tac (!claset addSDs [equalityD1]) 1);
by (subgoal_tac "xa: i" 1);
-by (fast_tac (ZF_cs addSDs [equalityD1]) 2);
-by (fast_tac (ZF_cs addSDs [equalityD1]
+by (fast_tac (!claset addSDs [equalityD1]) 2);
+by (fast_tac (!claset addSDs [equalityD1]
addSEs [bspec RS bspec RS bspec RS mp RS mp]) 1);
qed "Ord_alt_is_Ord";
@@ -318,27 +320,27 @@
goal OrderType.thy "(lam z:A+0. case(%x.x, %y.y, z)) : bij(A+0, A)";
by (res_inst_tac [("d", "Inl")] lam_bijective 1);
by (safe_tac sum_cs);
-by (ALLGOALS (asm_simp_tac sum_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "bij_sum_0";
goal OrderType.thy
"!!A r. well_ord(A,r) ==> ordertype(A+0, radd(A,r,0,s)) = ordertype(A,r)";
by (resolve_tac [bij_sum_0 RS ord_isoI RS ordertype_eq] 1);
by (assume_tac 2);
-by (fast_tac (sum_cs addss (sum_ss addsimps [radd_Inl_iff, Memrel_iff])) 1);
+by (fast_tac (sum_cs addss (!simpset addsimps [radd_Inl_iff, Memrel_iff])) 1);
qed "ordertype_sum_0_eq";
goal OrderType.thy "(lam z:0+A. case(%x.x, %y.y, z)) : bij(0+A, A)";
by (res_inst_tac [("d", "Inr")] lam_bijective 1);
by (safe_tac sum_cs);
-by (ALLGOALS (asm_simp_tac sum_ss));
+by (ALLGOALS (Asm_simp_tac));
qed "bij_0_sum";
goal OrderType.thy
"!!A r. well_ord(A,r) ==> ordertype(0+A, radd(0,s,A,r)) = ordertype(A,r)";
by (resolve_tac [bij_0_sum RS ord_isoI RS ordertype_eq] 1);
by (assume_tac 2);
-by (fast_tac (sum_cs addss (sum_ss addsimps [radd_Inr_iff, Memrel_iff])) 1);
+by (fast_tac (sum_cs addss (!simpset addsimps [radd_Inr_iff, Memrel_iff])) 1);
qed "ordertype_0_sum_eq";
(** Initial segments of radd. Statements by Grabczewski **)
@@ -352,7 +354,7 @@
by (safe_tac sum_cs);
by (ALLGOALS
(asm_full_simp_tac
- (sum_ss addsimps [radd_Inl_iff, radd_Inr_Inl_iff])));
+ (!simpset addsimps [radd_Inl_iff, radd_Inr_Inl_iff])));
qed "pred_Inl_bij";
goal OrderType.thy
@@ -361,7 +363,7 @@
\ ordertype(pred(A,a,r), r)";
by (resolve_tac [pred_Inl_bij RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1);
by (REPEAT_FIRST (ares_tac [pred_subset, well_ord_subset]));
-by (asm_full_simp_tac (ZF_ss addsimps [radd_Inl_iff, pred_def]) 1);
+by (asm_full_simp_tac (!simpset addsimps [radd_Inl_iff, pred_def]) 1);
qed "ordertype_pred_Inl_eq";
goalw OrderType.thy [pred_def, id_def]
@@ -370,7 +372,7 @@
\ : bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))";
by (res_inst_tac [("d", "%z.z")] lam_bijective 1);
by (safe_tac sum_cs);
-by (ALLGOALS (asm_full_simp_tac radd_ss));
+by (ALLGOALS (Asm_full_simp_tac));
qed "pred_Inr_bij";
goal OrderType.thy
@@ -378,7 +380,7 @@
\ ordertype(pred(A+B, Inr(b), radd(A,r,B,s)), radd(A,r,B,s)) = \
\ ordertype(A+pred(B,b,s), radd(A,r,pred(B,b,s),s))";
by (resolve_tac [pred_Inr_bij RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1);
-by (fast_tac (sum_cs addss (radd_ss addsimps [pred_def, id_def])) 2);
+by (fast_tac (sum_cs addss (!simpset addsimps [pred_def, id_def])) 2);
by (REPEAT_FIRST (ares_tac [well_ord_radd, pred_subset, well_ord_subset]));
qed "ordertype_pred_Inr_eq";
@@ -392,15 +394,16 @@
(** Ordinal addition with zero **)
goalw OrderType.thy [oadd_def] "!!i. Ord(i) ==> i++0 = i";
-by (asm_simp_tac (ZF_ss addsimps [Memrel_0, ordertype_sum_0_eq,
+by (asm_simp_tac (!simpset addsimps [Memrel_0, ordertype_sum_0_eq,
ordertype_Memrel, well_ord_Memrel]) 1);
qed "oadd_0";
goalw OrderType.thy [oadd_def] "!!i. Ord(i) ==> 0++i = i";
-by (asm_simp_tac (ZF_ss addsimps [Memrel_0, ordertype_0_sum_eq,
+by (asm_simp_tac (!simpset addsimps [Memrel_0, ordertype_0_sum_eq,
ordertype_Memrel, well_ord_Memrel]) 1);
qed "oadd_0_left";
+Addsimps [oadd_0, oadd_0_left];
(*** Further properties of ordinal addition. Statements by Grabczewski,
proofs by lcp. ***)
@@ -410,11 +413,11 @@
by (rtac ltI 1);
by (REPEAT (ares_tac [Ord_ordertype, well_ord_radd, well_ord_Memrel] 2));
by (asm_simp_tac
- (ZF_ss addsimps [ordertype_pred_unfold,
- well_ord_radd, well_ord_Memrel,
- ordertype_pred_Inl_eq,
- lt_pred_Memrel, leI RS le_ordertype_Memrel]
- setloop rtac (InlI RSN (2,RepFun_eqI))) 1);
+ (!simpset addsimps [ordertype_pred_unfold,
+ well_ord_radd, well_ord_Memrel,
+ ordertype_pred_Inl_eq,
+ lt_pred_Memrel, leI RS le_ordertype_Memrel]
+ setloop rtac (InlI RSN (2,bexI))) 1);
qed "lt_oadd1";
(*Thus also we obtain the rule i++j = k ==> i le k *)
@@ -428,8 +431,8 @@
goal OrderType.thy
"!!A B. A<=B ==> id(A) : ord_iso(A, Memrel(A), A, Memrel(B))";
by (resolve_tac [id_bij RS ord_isoI] 1);
-by (asm_simp_tac (ZF_ss addsimps [id_conv, Memrel_iff]) 1);
-by (fast_tac ZF_cs 1);
+by (asm_simp_tac (!simpset addsimps [id_conv, Memrel_iff]) 1);
+by (Fast_tac 1);
qed "id_ord_iso_Memrel";
goal OrderType.thy
@@ -449,7 +452,7 @@
by (rtac RepFun_eqI 1);
by (etac InrI 2);
by (asm_simp_tac
- (ZF_ss addsimps [ordertype_pred_Inr_eq, well_ord_Memrel,
+ (!simpset addsimps [ordertype_pred_Inr_eq, well_ord_Memrel,
lt_pred_Memrel, leI RS le_ordertype_Memrel,
ordertype_sum_Memrel]) 1);
qed "oadd_lt_mono2";
@@ -459,12 +462,12 @@
by (rtac Ord_linear_lt 1);
by (REPEAT_SOME assume_tac);
by (ALLGOALS
- (fast_tac (ZF_cs addDs [oadd_lt_mono2] addEs [lt_irrefl, lt_asym])));
+ (fast_tac (!claset addDs [oadd_lt_mono2] addEs [lt_irrefl, lt_asym])));
qed "oadd_lt_cancel2";
goal OrderType.thy
"!!i j k. [| Ord(i); Ord(j); Ord(k) |] ==> i++j < i++k <-> j<k";
-by (fast_tac (ZF_cs addSIs [oadd_lt_mono2] addSEs [oadd_lt_cancel2]) 1);
+by (fast_tac (!claset addSIs [oadd_lt_mono2] addSEs [oadd_lt_cancel2]) 1);
qed "oadd_lt_iff2";
goal OrderType.thy
@@ -472,8 +475,8 @@
by (rtac Ord_linear_lt 1);
by (REPEAT_SOME assume_tac);
by (ALLGOALS
- (fast_tac (ZF_cs addDs [oadd_lt_mono2]
- addss (ZF_ss addsimps [lt_not_refl]))));
+ (fast_tac (!claset addDs [oadd_lt_mono2]
+ addss (!simpset addsimps [lt_not_refl]))));
qed "oadd_inject";
goalw OrderType.thy [oadd_def]
@@ -481,11 +484,11 @@
(*Rotate the hypotheses so that simplification will work*)
by (etac revcut_rl 1);
by (asm_full_simp_tac
- (ZF_ss addsimps [ordertype_pred_unfold, well_ord_radd,
+ (!simpset addsimps [ordertype_pred_unfold, well_ord_radd,
well_ord_Memrel]) 1);
by (eresolve_tac [ltD RS RepFunE] 1);
by (fast_tac (sum_cs addss
- (ZF_ss addsimps [ordertype_pred_Inl_eq, well_ord_Memrel,
+ (!simpset addsimps [ordertype_pred_Inl_eq, well_ord_Memrel,
ltI, lt_pred_Memrel, le_ordertype_Memrel, leI,
ordertype_pred_Inr_eq,
ordertype_sum_Memrel])) 1);
@@ -511,21 +514,22 @@
by (rtac (subsetI RS equalityI) 1);
by (eresolve_tac [ltI RS lt_oadd_disj RS disjE] 1);
by (REPEAT (ares_tac [Ord_oadd] 1));
-by (fast_tac (ZF_cs addIs [lt_oadd1, oadd_lt_mono2]
- addss (ZF_ss addsimps [Ord_mem_iff_lt, Ord_oadd])) 3);
-by (fast_tac ZF_cs 2);
-by (fast_tac (ZF_cs addSEs [ltE]) 1);
+by (fast_tac (!claset addIs [lt_oadd1, oadd_lt_mono2]
+ addss (!simpset addsimps [Ord_mem_iff_lt, Ord_oadd])) 3);
+by (Fast_tac 2);
+by (fast_tac (!claset addSEs [ltE]) 1);
qed "oadd_unfold";
goal OrderType.thy "!!i. Ord(i) ==> i++1 = succ(i)";
-by (asm_simp_tac (ZF_ss addsimps [oadd_unfold, Ord_1, oadd_0]) 1);
+by (asm_simp_tac (!simpset addsimps [oadd_unfold, Ord_1, oadd_0]) 1);
by (fast_tac eq_cs 1);
qed "oadd_1";
goal OrderType.thy
"!!i. [| Ord(i); Ord(j) |] ==> i++succ(j) = succ(i++j)";
-by (asm_simp_tac
- (ZF_ss addsimps [oadd_1 RS sym, Ord_oadd, oadd_assoc, Ord_1]) 1);
+ (*ZF_ss prevents looping*)
+by (asm_simp_tac (ZF_ss addsimps [Ord_oadd, oadd_1 RS sym]) 1);
+by (asm_simp_tac (!simpset addsimps [oadd_1, oadd_assoc, Ord_1]) 1);
qed "oadd_succ";
@@ -542,7 +546,7 @@
goal OrderType.thy
"!!i j. [| Ord(i); Limit(j) |] ==> i++j = (UN k:j. i++k)";
by (forward_tac [Limit_has_0 RS ltD] 1);
-by (asm_simp_tac (ZF_ss addsimps [Limit_is_Ord RS Ord_in_Ord,
+by (asm_simp_tac (!simpset addsimps [Limit_is_Ord RS Ord_in_Ord,
oadd_UN RS sym, Union_eq_UN RS sym,
Limit_Union_eq]) 1);
qed "oadd_Limit";
@@ -551,25 +555,25 @@
goal OrderType.thy "!!i j. [| Ord(i); Ord(j) |] ==> i le j++i";
by (eres_inst_tac [("i","i")] trans_induct3 1);
-by (asm_simp_tac (ZF_ss addsimps [oadd_0, Ord_0_le]) 1);
-by (asm_simp_tac (ZF_ss addsimps [oadd_succ, succ_leI]) 1);
-by (asm_simp_tac (ZF_ss addsimps [oadd_Limit]) 1);
+by (asm_simp_tac (!simpset addsimps [Ord_0_le]) 1);
+by (asm_simp_tac (!simpset addsimps [oadd_succ, succ_leI]) 1);
+by (asm_simp_tac (!simpset addsimps [oadd_Limit]) 1);
by (rtac le_trans 1);
by (rtac le_implies_UN_le_UN 2);
-by (fast_tac ZF_cs 2);
-by (asm_simp_tac (ZF_ss addsimps [Union_eq_UN RS sym, Limit_Union_eq,
- le_refl, Limit_is_Ord]) 1);
+by (Fast_tac 2);
+by (asm_simp_tac (!simpset addsimps [Union_eq_UN RS sym, Limit_Union_eq,
+ le_refl, Limit_is_Ord]) 1);
qed "oadd_le_self2";
goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> k++i le j++i";
by (forward_tac [lt_Ord] 1);
by (forward_tac [le_Ord2] 1);
by (etac trans_induct3 1);
-by (asm_simp_tac (ZF_ss addsimps [oadd_0]) 1);
-by (asm_simp_tac (ZF_ss addsimps [oadd_succ, succ_le_iff]) 1);
-by (asm_simp_tac (ZF_ss addsimps [oadd_Limit]) 1);
+by (Asm_simp_tac 1);
+by (asm_simp_tac (!simpset addsimps [oadd_succ, succ_le_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [oadd_Limit]) 1);
by (rtac le_implies_UN_le_UN 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "oadd_le_mono1";
goal OrderType.thy "!!i j. [| i' le i; j'<j |] ==> i'++j' < i++j";
@@ -579,12 +583,12 @@
qed "oadd_lt_mono";
goal OrderType.thy "!!i j. [| i' le i; j' le j |] ==> i'++j' le i++j";
-by (asm_simp_tac (ZF_ss addsimps [oadd_succ RS sym, le_Ord2, oadd_lt_mono]) 1);
+by (asm_simp_tac (!simpset addsimps [oadd_succ RS sym, le_Ord2, oadd_lt_mono]) 1);
qed "oadd_le_mono";
goal OrderType.thy
"!!i j k. [| Ord(i); Ord(j); Ord(k) |] ==> i++j le i++k <-> j le k";
-by (asm_simp_tac (ZF_ss addsimps [oadd_lt_iff2, oadd_succ RS sym,
+by (asm_simp_tac (!simpset addsimps [oadd_lt_iff2, oadd_succ RS sym,
Ord_succ]) 1);
qed "oadd_le_iff2";
@@ -597,31 +601,31 @@
"!!A B. A<=B ==> (lam y:B. if(y:A, Inl(y), Inr(y))) : bij(B, A+(B-A))";
by (res_inst_tac [("d", "case(%x.x, %y.y)")] lam_bijective 1);
by (fast_tac (sum_cs addSIs [if_type]) 1);
-by (fast_tac (ZF_cs addSIs [case_type]) 1);
+by (fast_tac (!claset addSIs [case_type]) 1);
by (etac sumE 2);
-by (ALLGOALS (asm_simp_tac (sum_ss setloop split_tac [expand_if])));
+by (ALLGOALS (asm_simp_tac (!simpset setloop split_tac [expand_if])));
qed "bij_sum_Diff";
goal OrderType.thy
"!!i j. i le j ==> \
\ ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j))) = \
\ ordertype(j, Memrel(j))";
-by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1]));
+by (safe_tac (!claset addSDs [le_subset_iff RS iffD1]));
by (resolve_tac [bij_sum_Diff RS ord_isoI RS ord_iso_sym RS ordertype_eq] 1);
by (etac well_ord_Memrel 3);
by (assume_tac 1);
by (asm_simp_tac
- (radd_ss setloop split_tac [expand_if] addsimps [Memrel_iff]) 1);
+ (!simpset setloop split_tac [expand_if] addsimps [Memrel_iff]) 1);
by (forw_inst_tac [("j", "y")] Ord_in_Ord 1 THEN assume_tac 1);
by (forw_inst_tac [("j", "x")] Ord_in_Ord 1 THEN assume_tac 1);
-by (asm_simp_tac (ZF_ss addsimps [Ord_mem_iff_lt, lt_Ord, not_lt_iff_le]) 1);
-by (fast_tac (ZF_cs addEs [lt_trans2, lt_trans]) 1);
+by (asm_simp_tac (!simpset addsimps [Ord_mem_iff_lt, lt_Ord, not_lt_iff_le]) 1);
+by (fast_tac (!claset addEs [lt_trans2, lt_trans]) 1);
qed "ordertype_sum_Diff";
goalw OrderType.thy [oadd_def, odiff_def]
"!!i j. i le j ==> \
\ i ++ (j--i) = ordertype(i+(j-i), radd(i,Memrel(j),j-i,Memrel(j)))";
-by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1]));
+by (safe_tac (!claset addSDs [le_subset_iff RS iffD1]));
by (resolve_tac [sum_ord_iso_cong RS ordertype_eq] 1);
by (etac id_ord_iso_Memrel 1);
by (resolve_tac [ordertype_ord_iso RS ord_iso_sym] 1);
@@ -630,7 +634,7 @@
qed "oadd_ordertype_Diff";
goal OrderType.thy "!!i j. i le j ==> i ++ (j--i) = j";
-by (asm_simp_tac (ZF_ss addsimps [oadd_ordertype_Diff, ordertype_sum_Diff,
+by (asm_simp_tac (!simpset addsimps [oadd_ordertype_Diff, ordertype_sum_Diff,
ordertype_Memrel, lt_Ord2 RS Ord_succD]) 1);
qed "oadd_odiff_inverse";
@@ -646,14 +650,14 @@
"!!i j. [| Ord(i); Ord(j) |] ==> (i++j) -- i = j";
by (rtac oadd_inject 1);
by (REPEAT (ares_tac [Ord_ordertype, Ord_oadd, Ord_odiff] 2));
-by (asm_simp_tac (ZF_ss addsimps [oadd_odiff_inverse, oadd_le_self]) 1);
+by (asm_simp_tac (!simpset addsimps [oadd_odiff_inverse, oadd_le_self]) 1);
qed "odiff_oadd_inverse";
val [i_lt_j, k_le_i] = goal OrderType.thy
"[| i<j; k le i |] ==> i--k < j--k";
by (rtac (k_le_i RS lt_Ord RSN (2,oadd_lt_cancel2)) 1);
by (simp_tac
- (ZF_ss addsimps [i_lt_j, k_le_i, [k_le_i, leI] MRS le_trans,
+ (!simpset addsimps [i_lt_j, k_le_i, [k_le_i, leI] MRS le_trans,
oadd_odiff_inverse]) 1);
by (REPEAT (resolve_tac (Ord_odiff ::
([i_lt_j, k_le_i] RL [lt_Ord, lt_Ord2])) 1));
@@ -674,8 +678,8 @@
\ pred(A*B, <a,b>, rmult(A,r,B,s)) = \
\ pred(A,a,r)*B Un ({a} * pred(B,b,s))";
by (safe_tac eq_cs);
-by (ALLGOALS (asm_full_simp_tac (ZF_ss addsimps [rmult_iff])));
-by (ALLGOALS (fast_tac ZF_cs));
+by (ALLGOALS (asm_full_simp_tac (!simpset addsimps [rmult_iff])));
+by (ALLGOALS (Fast_tac));
qed "pred_Pair_eq";
goal OrderType.thy
@@ -683,11 +687,11 @@
\ ordertype(pred(A*B, <a,b>, rmult(A,r,B,s)), rmult(A,r,B,s)) = \
\ ordertype(pred(A,a,r)*B + pred(B,b,s), \
\ radd(A*B, rmult(A,r,B,s), B, s))";
-by (asm_simp_tac (ZF_ss addsimps [pred_Pair_eq]) 1);
+by (asm_simp_tac (!simpset addsimps [pred_Pair_eq]) 1);
by (resolve_tac [ordertype_eq RS sym] 1);
by (rtac prod_sum_singleton_ord_iso 1);
by (REPEAT_FIRST (ares_tac [pred_subset, well_ord_rmult RS well_ord_subset]));
-by (fast_tac (ZF_cs addSEs [predE]) 1);
+by (fast_tac (!claset addSEs [predE]) 1);
qed "ordertype_pred_Pair_eq";
goalw OrderType.thy [oadd_def, omult_def]
@@ -695,7 +699,7 @@
\ ordertype(pred(i*j, <i',j'>, rmult(i,Memrel(i),j,Memrel(j))), \
\ rmult(i,Memrel(i),j,Memrel(j))) = \
\ j**i' ++ j'";
-by (asm_simp_tac (ZF_ss addsimps [ordertype_pred_Pair_eq, lt_pred_Memrel,
+by (asm_simp_tac (!simpset addsimps [ordertype_pred_Pair_eq, lt_pred_Memrel,
ltD, lt_Ord2, well_ord_Memrel]) 1);
by (rtac trans 1);
by (resolve_tac [ordertype_ord_iso RS sum_ord_iso_cong RS ordertype_eq] 2);
@@ -705,35 +709,36 @@
by (REPEAT_FIRST (ares_tac [well_ord_rmult, well_ord_radd, well_ord_Memrel,
Ord_ordertype]));
by (ALLGOALS
- (asm_simp_tac (radd_ss addsimps [rmult_iff, id_conv, Memrel_iff])));
-by (safe_tac ZF_cs);
-by (ALLGOALS (fast_tac (ZF_cs addEs [Ord_trans])));
+ (asm_simp_tac (!simpset addsimps [rmult_iff, id_conv, Memrel_iff])));
+by (safe_tac (!claset));
+by (ALLGOALS (fast_tac (!claset addEs [Ord_trans])));
qed "ordertype_pred_Pair_lemma";
goalw OrderType.thy [omult_def]
"!!i j. [| Ord(i); Ord(j); k<j**i |] ==> \
\ EX j' i'. k = j**i' ++ j' & j'<j & i'<i";
-by (asm_full_simp_tac (ZF_ss addsimps [ordertype_pred_unfold,
+by (asm_full_simp_tac (!simpset addsimps [ordertype_pred_unfold,
well_ord_rmult, well_ord_Memrel]) 1);
-by (step_tac (ZF_cs addSEs [ltE]) 1);
-by (asm_simp_tac (ZF_ss addsimps [ordertype_pred_Pair_lemma, ltI,
+by (step_tac (!claset addSEs [ltE]) 1);
+by (asm_simp_tac (!simpset addsimps [ordertype_pred_Pair_lemma, ltI,
symmetric omult_def]) 1);
-by (fast_tac (ZF_cs addIs [ltI]) 1);
+by (fast_tac (!claset addIs [ltI]) 1);
qed "lt_omult";
goalw OrderType.thy [omult_def]
"!!i j. [| j'<j; i'<i |] ==> j**i' ++ j' < j**i";
by (rtac ltI 1);
by (asm_simp_tac
- (ZF_ss addsimps [Ord_ordertype, well_ord_rmult, well_ord_Memrel,
- lt_Ord2]) 2);
+ (!simpset addsimps [Ord_ordertype, well_ord_rmult, well_ord_Memrel,
+ lt_Ord2]) 2);
by (asm_simp_tac
- (ZF_ss addsimps [ordertype_pred_unfold,
+ (!simpset addsimps [ordertype_pred_unfold,
well_ord_rmult, well_ord_Memrel, lt_Ord2]) 1);
-by (rtac RepFun_eqI 1);
-by (fast_tac (ZF_cs addSEs [ltE]) 2);
+by (rtac bexI 1);
+by (fast_tac (!claset addSEs [ltE]) 2);
by (asm_simp_tac
- (ZF_ss addsimps [ordertype_pred_Pair_lemma, ltI, symmetric omult_def]) 1);
+ (!simpset addsimps [ordertype_pred_Pair_lemma, ltI,
+ symmetric omult_def]) 1);
qed "omult_oadd_lt";
goal OrderType.thy
@@ -742,8 +747,8 @@
by (resolve_tac [lt_omult RS exE] 1);
by (etac ltI 3);
by (REPEAT (ares_tac [Ord_omult] 1));
-by (fast_tac (ZF_cs addSEs [ltE]) 1);
-by (fast_tac (ZF_cs addIs [omult_oadd_lt RS ltD, ltI]) 1);
+by (fast_tac (!claset addSEs [ltE]) 1);
+by (fast_tac (!claset addIs [omult_oadd_lt RS ltD, ltI]) 1);
qed "omult_unfold";
(*** Basic laws for ordinal multiplication ***)
@@ -751,13 +756,15 @@
(** Ordinal multiplication by zero **)
goalw OrderType.thy [omult_def] "i**0 = 0";
-by (asm_simp_tac (ZF_ss addsimps [ordertype_0]) 1);
+by (Asm_simp_tac 1);
qed "omult_0";
goalw OrderType.thy [omult_def] "0**i = 0";
-by (asm_simp_tac (ZF_ss addsimps [ordertype_0]) 1);
+by (Asm_simp_tac 1);
qed "omult_0_left";
+Addsimps [omult_0, omult_0_left];
+
(** Ordinal multiplication by 1 **)
goalw OrderType.thy [omult_def] "!!i. Ord(i) ==> i**1 = i";
@@ -765,7 +772,7 @@
by (res_inst_tac [("c", "snd"), ("d", "%z.<0,z>")] lam_bijective 1);
by (REPEAT_FIRST (eresolve_tac [snd_type, SigmaE, succE, emptyE,
well_ord_Memrel, ordertype_Memrel]));
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [rmult_iff, Memrel_iff])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [rmult_iff, Memrel_iff])));
qed "omult_1";
goalw OrderType.thy [omult_def] "!!i. Ord(i) ==> 1**i = i";
@@ -773,9 +780,11 @@
by (res_inst_tac [("c", "fst"), ("d", "%z.<z,0>")] lam_bijective 1);
by (REPEAT_FIRST (eresolve_tac [fst_type, SigmaE, succE, emptyE,
well_ord_Memrel, ordertype_Memrel]));
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [rmult_iff, Memrel_iff])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [rmult_iff, Memrel_iff])));
qed "omult_1_left";
+Addsimps [omult_1, omult_1_left];
+
(** Distributive law for ordinal multiplication and addition **)
goalw OrderType.thy [omult_def, oadd_def]
@@ -793,8 +802,10 @@
qed "oadd_omult_distrib";
goal OrderType.thy "!!i. [| Ord(i); Ord(j) |] ==> i**succ(j) = (i**j)++i";
+ (*ZF_ss prevents looping*)
+by (asm_simp_tac (ZF_ss addsimps [oadd_1 RS sym]) 1);
by (asm_simp_tac
- (ZF_ss addsimps [oadd_1 RS sym, omult_1, oadd_omult_distrib, Ord_1]) 1);
+ (!simpset addsimps [omult_1, oadd_omult_distrib, Ord_1]) 1);
qed "omult_succ";
(** Associative law **)
@@ -818,14 +829,14 @@
val prems = goal OrderType.thy
"[| Ord(i); !!x. x:A ==> Ord(j(x)) |] ==> \
\ i ** (UN x:A. j(x)) = (UN x:A. i**j(x))";
-by (asm_simp_tac (ZF_ss addsimps (prems@[Ord_UN, omult_unfold])) 1);
+by (asm_simp_tac (!simpset addsimps (prems@[Ord_UN, omult_unfold])) 1);
by (fast_tac eq_cs 1);
qed "omult_UN";
goal OrderType.thy
"!!i j. [| Ord(i); Limit(j) |] ==> i**j = (UN k:j. i**k)";
by (asm_simp_tac
- (ZF_ss addsimps [Limit_is_Ord RS Ord_in_Ord, omult_UN RS sym,
+ (!simpset addsimps [Limit_is_Ord RS Ord_in_Ord, omult_UN RS sym,
Union_eq_UN RS sym, Limit_Union_eq]) 1);
qed "omult_Limit";
@@ -834,10 +845,10 @@
(*As a special case we have "[| 0<i; 0<j |] ==> 0 < i**j" *)
goal OrderType.thy "!!i j. [| k<i; 0<j |] ==> k < i**j";
-by (safe_tac (ZF_cs addSEs [ltE] addSIs [ltI, Ord_omult]));
-by (asm_simp_tac (ZF_ss addsimps [omult_unfold]) 1);
-by (REPEAT (etac UN_I 1));
-by (asm_simp_tac (ZF_ss addsimps [omult_0, oadd_0_left]) 1);
+by (safe_tac (!claset addSEs [ltE] addSIs [ltI, Ord_omult]));
+by (asm_simp_tac (!simpset addsimps [omult_unfold]) 1);
+by (REPEAT_FIRST (ares_tac [bexI]));
+by (Asm_simp_tac 1);
qed "lt_omult1";
goal OrderType.thy "!!i j. [| Ord(i); 0<j |] ==> i le i**j";
@@ -849,26 +860,26 @@
by (forward_tac [lt_Ord] 1);
by (forward_tac [le_Ord2] 1);
by (etac trans_induct3 1);
-by (asm_simp_tac (ZF_ss addsimps [omult_0, le_refl, Ord_0]) 1);
-by (asm_simp_tac (ZF_ss addsimps [omult_succ, oadd_le_mono]) 1);
-by (asm_simp_tac (ZF_ss addsimps [omult_Limit]) 1);
+by (asm_simp_tac (!simpset addsimps [le_refl, Ord_0]) 1);
+by (asm_simp_tac (!simpset addsimps [omult_succ, oadd_le_mono]) 1);
+by (asm_simp_tac (!simpset addsimps [omult_Limit]) 1);
by (rtac le_implies_UN_le_UN 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "omult_le_mono1";
goal OrderType.thy "!!i j k. [| k<j; 0<i |] ==> i**k < i**j";
by (rtac ltI 1);
-by (asm_simp_tac (ZF_ss addsimps [omult_unfold, lt_Ord2]) 1);
-by (safe_tac (ZF_cs addSEs [ltE] addSIs [Ord_omult]));
-by (REPEAT (etac UN_I 1));
-by (asm_simp_tac (ZF_ss addsimps [omult_0, oadd_0, Ord_omult]) 1);
+by (asm_simp_tac (!simpset addsimps [omult_unfold, lt_Ord2]) 1);
+by (safe_tac (!claset addSEs [ltE] addSIs [Ord_omult]));
+by (REPEAT_FIRST (ares_tac [bexI]));
+by (asm_simp_tac (!simpset addsimps [Ord_omult]) 1);
qed "omult_lt_mono2";
goal OrderType.thy "!!i j k. [| k le j; Ord(i) |] ==> i**k le i**j";
by (rtac subset_imp_le 1);
-by (safe_tac (ZF_cs addSEs [ltE, make_elim Ord_succD] addSIs [Ord_omult]));
-by (asm_full_simp_tac (ZF_ss addsimps [omult_unfold]) 1);
-by (deepen_tac (ZF_cs addEs [Ord_trans, UN_I]) 0 1);
+by (safe_tac (!claset addSEs [ltE, make_elim Ord_succD] addSIs [Ord_omult]));
+by (asm_full_simp_tac (!simpset addsimps [omult_unfold]) 1);
+by (deepen_tac (!claset addEs [Ord_trans]) 0 1);
qed "omult_le_mono2";
goal OrderType.thy "!!i j. [| i' le i; j' le j |] ==> i'**j' le i**j";
@@ -887,18 +898,18 @@
goal OrderType.thy "!!i j. [| Ord(i); 0<j |] ==> i le j**i";
by (forward_tac [lt_Ord2] 1);
by (eres_inst_tac [("i","i")] trans_induct3 1);
-by (asm_simp_tac (ZF_ss addsimps [omult_0, Ord_0 RS le_refl]) 1);
-by (asm_simp_tac (ZF_ss addsimps [omult_succ, succ_le_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [omult_0, Ord_0 RS le_refl]) 1);
+by (asm_simp_tac (!simpset addsimps [omult_succ, succ_le_iff]) 1);
by (etac lt_trans1 1);
by (res_inst_tac [("b", "j**x")] (oadd_0 RS subst) 1 THEN
rtac oadd_lt_mono2 2);
by (REPEAT (ares_tac [Ord_omult] 1));
-by (asm_simp_tac (ZF_ss addsimps [omult_Limit]) 1);
+by (asm_simp_tac (!simpset addsimps [omult_Limit]) 1);
by (rtac le_trans 1);
by (rtac le_implies_UN_le_UN 2);
-by (fast_tac ZF_cs 2);
-by (asm_simp_tac (ZF_ss addsimps [Union_eq_UN RS sym, Limit_Union_eq,
- Limit_is_Ord RS le_refl]) 1);
+by (Fast_tac 2);
+by (asm_simp_tac (!simpset addsimps [Union_eq_UN RS sym, Limit_Union_eq,
+ Limit_is_Ord RS le_refl]) 1);
qed "omult_le_self2";
@@ -908,8 +919,8 @@
by (rtac Ord_linear_lt 1);
by (REPEAT_SOME assume_tac);
by (ALLGOALS
- (fast_tac (ZF_cs addDs [omult_lt_mono2]
- addss (ZF_ss addsimps [lt_not_refl]))));
+ (best_tac (!claset addDs [omult_lt_mono2]
+ addss (!simpset addsimps [lt_not_refl]))));
qed "omult_inject";
--- a/src/ZF/OrderType.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/OrderType.thy Fri Jan 03 15:01:55 1997 +0100
@@ -8,7 +8,7 @@
The order type of a well-ordering is the least ordinal isomorphic to it.
*)
-OrderType = OrderArith + Ordinal +
+OrderType = OrderArith + OrdQuant +
consts
ordermap :: [i,i]=>i
ordertype :: [i,i]=>i
--- a/src/ZF/Ordinal.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Ordinal.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,7 +13,7 @@
(** Two neat characterisations of Transset **)
goalw Ordinal.thy [Transset_def] "Transset(A) <-> A<=Pow(A)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_iff_Pow";
goalw Ordinal.thy [Transset_def] "Transset(A) <-> Union(succ(A)) = A";
@@ -24,64 +24,64 @@
goalw Ordinal.thy [Transset_def]
"!!C a b. [| Transset(C); {a,b}: C |] ==> a:C & b: C";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_doubleton_D";
val [prem1,prem2] = goalw Ordinal.thy [Pair_def]
"[| Transset(C); <a,b>: C |] ==> a:C & b: C";
by (cut_facts_tac [prem2] 1);
-by (fast_tac (ZF_cs addSDs [prem1 RS Transset_doubleton_D]) 1);
+by (fast_tac (!claset addSDs [prem1 RS Transset_doubleton_D]) 1);
qed "Transset_Pair_D";
val prem1::prems = goal Ordinal.thy
"[| Transset(C); A*B <= C; b: B |] ==> A <= C";
by (cut_facts_tac prems 1);
-by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1);
+by (fast_tac (!claset addSDs [prem1 RS Transset_Pair_D]) 1);
qed "Transset_includes_domain";
val prem1::prems = goal Ordinal.thy
"[| Transset(C); A*B <= C; a: A |] ==> B <= C";
by (cut_facts_tac prems 1);
-by (fast_tac (ZF_cs addSDs [prem1 RS Transset_Pair_D]) 1);
+by (fast_tac (!claset addSDs [prem1 RS Transset_Pair_D]) 1);
qed "Transset_includes_range";
(** Closure properties **)
goalw Ordinal.thy [Transset_def] "Transset(0)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_0";
goalw Ordinal.thy [Transset_def]
"!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Un j)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_Un";
goalw Ordinal.thy [Transset_def]
"!!i j. [| Transset(i); Transset(j) |] ==> Transset(i Int j)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_Int";
goalw Ordinal.thy [Transset_def] "!!i. Transset(i) ==> Transset(succ(i))";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_succ";
goalw Ordinal.thy [Transset_def] "!!i. Transset(i) ==> Transset(Pow(i))";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_Pow";
goalw Ordinal.thy [Transset_def] "!!A. Transset(A) ==> Transset(Union(A))";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_Union";
val [Transprem] = goalw Ordinal.thy [Transset_def]
"[| !!i. i:A ==> Transset(i) |] ==> Transset(Union(A))";
-by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1);
+by (fast_tac (!claset addEs [Transprem RS bspec RS subsetD]) 1);
qed "Transset_Union_family";
val [prem,Transprem] = goalw Ordinal.thy [Transset_def]
"[| j:A; !!i. i:A ==> Transset(i) |] ==> Transset(Inter(A))";
by (cut_facts_tac [prem] 1);
-by (fast_tac (ZF_cs addEs [Transprem RS bspec RS subsetD]) 1);
+by (fast_tac (!claset addEs [Transprem RS bspec RS subsetD]) 1);
qed "Transset_Inter_family";
(*** Natural Deduction rules for Ord ***)
@@ -104,7 +104,7 @@
(*** Lemmas for ordinals ***)
goalw Ordinal.thy [Ord_def,Transset_def] "!!i j.[| Ord(i); j:i |] ==> Ord(j)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Ord_in_Ord";
(* Ord(succ(j)) ==> Ord(j) *)
@@ -116,7 +116,7 @@
qed "Ord_subset_Ord";
goalw Ordinal.thy [Ord_def,Transset_def] "!!i j. [| j:i; Ord(i) |] ==> j<=i";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "OrdmemD";
goal Ordinal.thy "!!i j k. [| i:j; j:k; Ord(k) |] ==> i:k";
@@ -143,15 +143,18 @@
bind_thm ("Ord_1", Ord_0 RS Ord_succ);
goal Ordinal.thy "Ord(succ(i)) <-> Ord(i)";
-by (fast_tac (ZF_cs addIs [Ord_succ] addDs [Ord_succD]) 1);
+by (fast_tac (!claset addIs [Ord_succ] addDs [Ord_succD]) 1);
qed "Ord_succ_iff";
+Addsimps [Ord_0, Ord_succ_iff];
+AddSIs [Ord_0, Ord_succ];
+
goalw Ordinal.thy [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Un j)";
-by (fast_tac (ZF_cs addSIs [Transset_Un]) 1);
+by (fast_tac (!claset addSIs [Transset_Un]) 1);
qed "Ord_Un";
goalw Ordinal.thy [Ord_def] "!!i j. [| Ord(i); Ord(j) |] ==> Ord(i Int j)";
-by (fast_tac (ZF_cs addSIs [Transset_Int]) 1);
+by (fast_tac (!claset addSIs [Transset_Int]) 1);
qed "Ord_Int";
val nonempty::prems = goal Ordinal.thy
@@ -174,12 +177,12 @@
goal Ordinal.thy "~ (ALL i. i:X <-> Ord(i))";
by (rtac notI 1);
by (forw_inst_tac [("x", "X")] spec 1);
-by (safe_tac (ZF_cs addSEs [mem_irrefl]));
+by (safe_tac (!claset addSEs [mem_irrefl]));
by (swap_res_tac [Ord_is_Transset RSN (2,OrdI)] 1);
-by (fast_tac ZF_cs 2);
+by (Fast_tac 2);
by (rewtac Transset_def);
-by (safe_tac ZF_cs);
-by (asm_full_simp_tac ZF_ss 1);
+by (safe_tac (!claset));
+by (Asm_full_simp_tac 1);
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 1));
qed "ON_class";
@@ -201,9 +204,11 @@
qed "ltD";
goalw Ordinal.thy [lt_def] "~ i<0";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "not_lt0";
+Addsimps [not_lt0];
+
goal Ordinal.thy "!!i j. j<i ==> Ord(j)";
by (etac ltE 1 THEN assume_tac 1);
qed "lt_Ord";
@@ -219,7 +224,7 @@
bind_thm ("lt0E", not_lt0 RS notE);
goal Ordinal.thy "!!i j k. [| i<j; j<k |] ==> i<k";
-by (fast_tac (ZF_cs addSIs [ltI] addSEs [ltE, Ord_trans]) 1);
+by (fast_tac (!claset addSIs [ltI] addSEs [ltE, Ord_trans]) 1);
qed "lt_trans";
goalw Ordinal.thy [lt_def] "!!i j. [| i<j; j<i |] ==> P";
@@ -232,19 +237,21 @@
qed_goal "lt_not_refl" Ordinal.thy "~ i<i"
(fn _=> [ (rtac notI 1), (etac lt_irrefl 1) ]);
+AddSEs [lt_irrefl, lt0E];
+
(** le is less than or equals; recall i le j abbrevs i<succ(j) !! **)
goalw Ordinal.thy [lt_def] "i le j <-> i<j | (i=j & Ord(j))";
-by (fast_tac (ZF_cs addSIs [Ord_succ] addSDs [Ord_succD]) 1);
+by (fast_tac (!claset addSIs [Ord_succ] addSDs [Ord_succD]) 1);
qed "le_iff";
(*Equivalently, i<j ==> i < succ(j)*)
goal Ordinal.thy "!!i j. i<j ==> i le j";
-by (asm_simp_tac (ZF_ss addsimps [le_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [le_iff]) 1);
qed "leI";
goal Ordinal.thy "!!i. [| i=j; Ord(j) |] ==> i le j";
-by (asm_simp_tac (ZF_ss addsimps [le_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [le_iff]) 1);
qed "le_eqI";
val le_refl = refl RS le_eqI;
@@ -261,26 +268,27 @@
qed "leE";
goal Ordinal.thy "!!i j. [| i le j; j le i |] ==> i=j";
-by (asm_full_simp_tac (ZF_ss addsimps [le_iff]) 1);
-by (fast_tac (ZF_cs addEs [lt_asym]) 1);
+by (asm_full_simp_tac (!simpset addsimps [le_iff]) 1);
+by (fast_tac (!claset addEs [lt_asym]) 1);
qed "le_anti_sym";
goal Ordinal.thy "i le 0 <-> i=0";
-by (fast_tac (ZF_cs addSIs [Ord_0 RS le_refl] addSEs [leE, lt0E]) 1);
+by (fast_tac (!claset addSIs [Ord_0 RS le_refl] addSEs [leE]) 1);
qed "le0_iff";
bind_thm ("le0D", le0_iff RS iffD1);
-val lt_cs =
- ZF_cs addSIs [le_refl, leCI]
- addSDs [le0D]
- addSEs [lt_irrefl, lt0E, leE];
+AddIs [le_refl];
+AddSDs [le0D];
+Addsimps [le0_iff];
+
+val le_cs = !claset addSIs [leCI] addSEs [leE] addEs [lt_asym];
(*** Natural Deduction rules for Memrel ***)
goalw Ordinal.thy [Memrel_def] "<a,b> : Memrel(A) <-> a:b & a:A & b:A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Memrel_iff";
val prems = goal Ordinal.thy "[| a: b; a: A; b: A |] ==> <a,b> : Memrel(A)";
@@ -298,11 +306,11 @@
qed "MemrelE";
goalw Ordinal.thy [Memrel_def] "Memrel(A) <= A*A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Memrel_type";
goalw Ordinal.thy [Memrel_def] "!!A B. A<=B ==> Memrel(A) <= Memrel(B)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Memrel_mono";
goalw Ordinal.thy [Memrel_def] "Memrel(0) = 0";
@@ -313,6 +321,8 @@
by (fast_tac eq_cs 1);
qed "Memrel_1";
+Addsimps [Memrel_0, Memrel_1];
+
(*The membership relation (as a set) is well-founded.
Proof idea: show A<=B by applying the foundation axiom to A-B *)
goalw Ordinal.thy [wf_def] "wf(Memrel(A))";
@@ -328,13 +338,13 @@
(*Transset(i) does not suffice, though ALL j:i.Transset(j) does*)
goalw Ordinal.thy [Ord_def, Transset_def, trans_def]
"!!i. Ord(i) ==> trans(Memrel(i))";
-by (fast_tac (ZF_cs addSIs [MemrelI] addSEs [MemrelE]) 1);
+by (fast_tac (!claset addSIs [MemrelI] addSEs [MemrelE]) 1);
qed "trans_Memrel";
(*If Transset(A) then Memrel(A) internalizes the membership relation below A*)
goalw Ordinal.thy [Transset_def]
"!!A. Transset(A) ==> <a,b> : Memrel(A) <-> a:b & b:A";
-by (fast_tac (ZF_cs addSIs [MemrelI] addSEs [MemrelE]) 1);
+by (fast_tac (!claset addSIs [MemrelI] addSEs [MemrelE]) 1);
qed "Transset_Memrel_iff";
@@ -346,11 +356,11 @@
\ !!x.[| x: k; ALL y:x. P(y) |] ==> P(x) \
\ |] ==> P(i)";
by (rtac (major RS (wf_Memrel RS wf_induct2)) 1);
-by (fast_tac (ZF_cs addEs [MemrelE]) 1);
+by (fast_tac (!claset addEs [MemrelE]) 1);
by (resolve_tac prems 1);
by (assume_tac 1);
by (cut_facts_tac prems 1);
-by (fast_tac (ZF_cs addIs [MemrelI]) 1);
+by (fast_tac (!claset addIs [MemrelI]) 1);
qed "Transset_induct";
(*Induction over an ordinal*)
@@ -413,19 +423,19 @@
qed "Ord_linear_le";
goal Ordinal.thy "!!i j. j le i ==> ~ i<j";
-by (fast_tac (lt_cs addEs [lt_asym]) 1);
+by (fast_tac le_cs 1);
qed "le_imp_not_lt";
goal Ordinal.thy "!!i j. [| ~ i<j; Ord(i); Ord(j) |] ==> j le i";
by (res_inst_tac [("i","i"),("j","j")] Ord_linear2 1);
by (REPEAT (SOMEGOAL assume_tac));
-by (fast_tac (lt_cs addEs [lt_asym]) 1);
+by (fast_tac le_cs 1);
qed "not_lt_imp_le";
(** Some rewrite rules for <, le **)
goalw Ordinal.thy [lt_def] "!!i j. Ord(j) ==> i:j <-> i<j";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Ord_mem_iff_lt";
goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> ~ i<j <-> j le i";
@@ -433,7 +443,7 @@
qed "not_lt_iff_le";
goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> ~ i le j <-> j<i";
-by (asm_simp_tac (ZF_ss addsimps [not_lt_iff_le RS iff_sym]) 1);
+by (asm_simp_tac (!simpset addsimps [not_lt_iff_le RS iff_sym]) 1);
qed "not_le_iff_lt";
(*This is identical to 0<succ(i) *)
@@ -445,7 +455,7 @@
goal Ordinal.thy "!!i. [| Ord(i); i~=0 |] ==> 0<i";
by (etac (not_le_iff_lt RS iffD1) 1);
by (rtac Ord_0 1);
-by (fast_tac lt_cs 1);
+by (Fast_tac 1);
qed "Ord_0_lt";
(*** Results about less-than or equals ***)
@@ -456,23 +466,23 @@
by (rtac (not_lt_iff_le RS iffD1) 1);
by (assume_tac 1);
by (assume_tac 1);
-by (fast_tac (ZF_cs addEs [ltE, mem_irrefl]) 1);
+by (fast_tac (!claset addEs [ltE, mem_irrefl]) 1);
qed "subset_imp_le";
goal Ordinal.thy "!!i j. i le j ==> i<=j";
by (etac leE 1);
-by (fast_tac ZF_cs 2);
+by (Fast_tac 2);
by (fast_tac (subset_cs addIs [OrdmemD] addEs [ltE]) 1);
qed "le_imp_subset";
goal Ordinal.thy "j le i <-> j<=i & Ord(i) & Ord(j)";
-by (fast_tac (ZF_cs addSEs [subset_imp_le, le_imp_subset]
+by (fast_tac (!claset addSEs [subset_imp_le, le_imp_subset]
addEs [ltE, make_elim Ord_succD]) 1);
qed "le_subset_iff";
goal Ordinal.thy "i le succ(j) <-> i le j | i=succ(j) & Ord(i)";
-by (simp_tac (ZF_ss addsimps [le_iff]) 1);
-by (fast_tac (ZF_cs addIs [Ord_succ] addDs [Ord_succD]) 1);
+by (simp_tac (!simpset addsimps [le_iff]) 1);
+by (fast_tac (!claset addIs [Ord_succ] addDs [Ord_succD]) 1);
qed "le_succ_iff";
(*Just a variant of subset_imp_le*)
@@ -485,11 +495,11 @@
(** Transitive laws **)
goal Ordinal.thy "!!i j. [| i le j; j<k |] ==> i<k";
-by (fast_tac (ZF_cs addEs [leE, lt_trans]) 1);
+by (fast_tac (!claset addEs [leE, lt_trans]) 1);
qed "lt_trans1";
goal Ordinal.thy "!!i j. [| i<j; j le k |] ==> i<k";
-by (fast_tac (ZF_cs addEs [leE, lt_trans]) 1);
+by (fast_tac (!claset addEs [leE, lt_trans]) 1);
qed "lt_trans2";
goal Ordinal.thy "!!i j. [| i le j; j le k |] ==> i le k";
@@ -498,23 +508,25 @@
goal Ordinal.thy "!!i j. i<j ==> succ(i) le j";
by (rtac (not_lt_iff_le RS iffD1) 1);
-by (fast_tac (lt_cs addEs [lt_asym]) 3);
-by (ALLGOALS (fast_tac (ZF_cs addEs [ltE] addIs [Ord_succ])));
+by (fast_tac le_cs 3);
+by (ALLGOALS (fast_tac (!claset addEs [ltE] addIs [Ord_succ])));
qed "succ_leI";
(*Identical to succ(i) < succ(j) ==> i<j *)
goal Ordinal.thy "!!i j. succ(i) le j ==> i<j";
by (rtac (not_le_iff_lt RS iffD1) 1);
-by (fast_tac (lt_cs addEs [lt_asym]) 3);
-by (ALLGOALS (fast_tac (ZF_cs addEs [ltE, make_elim Ord_succD])));
+by (fast_tac le_cs 3);
+by (ALLGOALS (fast_tac (!claset addEs [ltE, make_elim Ord_succD])));
qed "succ_leE";
goal Ordinal.thy "succ(i) le j <-> i<j";
by (REPEAT (ares_tac [iffI,succ_leI,succ_leE] 1));
qed "succ_le_iff";
+Addsimps [succ_le_iff];
+
goal Ordinal.thy "!!i j. succ(i) le succ(j) ==> i le j";
-by (fast_tac (lt_cs addSEs [succ_leE]) 1);
+by (fast_tac (!claset addSEs [succ_leE]) 1);
qed "succ_le_imp_le";
(** Union and Intersection **)
@@ -533,13 +545,13 @@
goal Ordinal.thy "!!i j k. [| i<k; j<k |] ==> i Un j < k";
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1);
by (stac Un_commute 4);
-by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Un_iff]) 4);
-by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Un_iff]) 3);
+by (asm_full_simp_tac (!simpset addsimps [le_subset_iff, subset_Un_iff]) 4);
+by (asm_full_simp_tac (!simpset addsimps [le_subset_iff, subset_Un_iff]) 3);
by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
qed "Un_least_lt";
goal Ordinal.thy "!!i j. [| Ord(i); Ord(j) |] ==> i Un j < k <-> i<k & j<k";
-by (safe_tac (ZF_cs addSIs [Un_least_lt]));
+by (safe_tac (!claset addSIs [Un_least_lt]));
by (rtac (Un_upper2_le RS lt_trans1) 2);
by (rtac (Un_upper1_le RS lt_trans1) 1);
by (REPEAT_SOME assume_tac);
@@ -549,15 +561,15 @@
"[| Ord(i); Ord(j); Ord(k) |] ==> i Un j : k <-> i:k & j:k";
by (cut_facts_tac [[ordi,ordj] MRS
read_instantiate [("k","k")] Un_least_lt_iff] 1);
-by (asm_full_simp_tac (ZF_ss addsimps [lt_def,ordi,ordj,ordk]) 1);
+by (asm_full_simp_tac (!simpset addsimps [lt_def,ordi,ordj,ordk]) 1);
qed "Un_least_mem_iff";
(*Replacing k by succ(k') yields the similar rule for le!*)
goal Ordinal.thy "!!i j k. [| i<k; j<k |] ==> i Int j < k";
by (res_inst_tac [("i","i"),("j","j")] Ord_linear_le 1);
by (stac Int_commute 4);
-by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Int_iff]) 4);
-by (asm_full_simp_tac (ZF_ss addsimps [le_subset_iff, subset_Int_iff]) 3);
+by (asm_full_simp_tac (!simpset addsimps [le_subset_iff, subset_Int_iff]) 4);
+by (asm_full_simp_tac (!simpset addsimps [le_subset_iff, subset_Int_iff]) 3);
by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
qed "Int_greatest_lt";
@@ -616,7 +628,7 @@
(*Holds for all transitive sets, not just ordinals*)
goal Ordinal.thy "!!i. Ord(i) ==> Union(i) <= i";
-by (fast_tac (ZF_cs addSEs [Ord_trans]) 1);
+by (fast_tac (!claset addSEs [Ord_trans]) 1);
qed "Ord_Union_subset";
@@ -635,14 +647,14 @@
qed "Limit_has_0";
goalw Ordinal.thy [Limit_def] "!!i. [| Limit(i); j<i |] ==> succ(j) < i";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Limit_has_succ";
goalw Ordinal.thy [Limit_def]
"!!i. [| 0<i; ALL y. succ(y) ~= i |] ==> Limit(i)";
by (safe_tac subset_cs);
by (rtac (not_le_iff_lt RS iffD1) 2);
-by (fast_tac (lt_cs addEs [lt_asym]) 4);
+by (fast_tac le_cs 4);
by (REPEAT (eresolve_tac [asm_rl, ltE, Ord_succ] 1));
qed "non_succ_LimitI";
@@ -654,14 +666,14 @@
qed "succ_LimitE";
goal Ordinal.thy "!!i. [| Limit(i); i le succ(j) |] ==> i le j";
-by (safe_tac (ZF_cs addSEs [succ_LimitE, leE]));
+by (safe_tac (!claset addSEs [succ_LimitE, leE]));
qed "Limit_le_succD";
(** Traditional 3-way case analysis on ordinals **)
goal Ordinal.thy "!!i. Ord(i) ==> i=0 | (EX j. Ord(j) & i=succ(j)) | Limit(i)";
-by (fast_tac (ZF_cs addSIs [non_succ_LimitI, Ord_0_lt]
- addSDs [Ord_succD]) 1);
+by (fast_tac (!claset addSIs [non_succ_LimitI, Ord_0_lt]
+ addSDs [Ord_succD]) 1);
qed "Ord_cases_disj";
val major::prems = goal Ordinal.thy
@@ -682,5 +694,5 @@
\ |] ==> P(i)";
by (resolve_tac [major RS trans_induct] 1);
by (etac Ord_cases 1);
-by (ALLGOALS (fast_tac (ZF_cs addIs prems)));
+by (ALLGOALS (fast_tac (!claset addIs prems)));
qed "trans_induct3";
--- a/src/ZF/Ordinal.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Ordinal.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
Ordinals in Zermelo-Fraenkel Set Theory
*)
-Ordinal = WF + Bool + "simpdata" + "equalities" +
+Ordinal = WF + Bool + equalities +
consts
Memrel :: i=>i
Transset,Ord :: i=>o
--- a/src/ZF/Perm.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Perm.ML Fri Jan 03 15:01:55 1997 +0100
@@ -18,12 +18,12 @@
qed "surj_is_fun";
goalw Perm.thy [surj_def] "!!f A B. f : Pi(A,B) ==> f: surj(A,range(f))";
-by (fast_tac (ZF_cs addIs [apply_equality]
+by (fast_tac (!claset addIs [apply_equality]
addEs [range_of_fun,domain_type]) 1);
qed "fun_is_surj";
goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> range(f)=B";
-by (best_tac (ZF_cs addIs [equalityI,apply_Pair] addEs [range_type]) 1);
+by (best_tac (!claset addIs [equalityI,apply_Pair] addEs [range_type]) 1);
qed "surj_range";
(** A function with a right inverse is a surjection **)
@@ -31,7 +31,7 @@
val prems = goalw Perm.thy [surj_def]
"[| f: A->B; !!y. y:B ==> d(y): A; !!y. y:B ==> f`d(y) = y \
\ |] ==> f: surj(A,B)";
-by (fast_tac (ZF_cs addIs prems) 1);
+by (fast_tac (!claset addIs prems) 1);
qed "f_imp_surjective";
val prems = goal Perm.thy
@@ -40,12 +40,12 @@
\ !!y. y:B ==> c(d(y)) = y \
\ |] ==> (lam x:A.c(x)) : surj(A,B)";
by (res_inst_tac [("d", "d")] f_imp_surjective 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([lam_type]@prems)) ));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps ([lam_type]@prems)) ));
qed "lam_surjective";
(*Cantor's theorem revisited*)
goalw Perm.thy [surj_def] "f ~: surj(A,Pow(A))";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (cut_facts_tac [cantor] 1);
by (fast_tac subset_cs 1);
qed "cantor_surj";
@@ -61,18 +61,18 @@
goalw Perm.thy [inj_def]
"!!f A B. [| <a,b>:f; <c,b>:f; f: inj(A,B) |] ==> a=c";
by (REPEAT (eresolve_tac [asm_rl, Pair_mem_PiE, CollectE] 1));
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "inj_equality";
goalw thy [inj_def] "!!A B f. [| f:inj(A,B); a:A; b:A; f`a=f`b |] ==> a=b";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val inj_apply_equality = result();
(** A function with a left inverse is an injection **)
goal Perm.thy "!!f. [| f: A->B; ALL x:A. d(f`x)=x |] ==> f: inj(A,B)";
-by (asm_simp_tac (ZF_ss addsimps [inj_def]) 1);
-by (deepen_tac (ZF_cs addEs [subst_context RS box_equals]) 0 1);
+by (asm_simp_tac (!simpset addsimps [inj_def]) 1);
+by (deepen_tac (!claset addEs [subst_context RS box_equals]) 0 1);
bind_thm ("f_imp_injective", ballI RSN (2,result()));
val prems = goal Perm.thy
@@ -80,7 +80,7 @@
\ !!x. x:A ==> d(c(x)) = x \
\ |] ==> (lam x:A.c(x)) : inj(A,B)";
by (res_inst_tac [("d", "d")] f_imp_injective 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([lam_type]@prems)) ));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps ([lam_type]@prems)) ));
qed "lam_injective";
(** Bijections **)
@@ -125,8 +125,10 @@
qed "id_type";
goalw Perm.thy [id_def] "!!A x. x:A ==> id(A)`x = x";
-by (asm_simp_tac ZF_ss 1);
-val id_conv = result();
+by (Asm_simp_tac 1);
+qed "id_conv";
+
+Addsimps [id_conv];
val [prem] = goalw Perm.thy [id_def] "A<=B ==> id(A) <= id(B)";
by (rtac (prem RS lam_mono) 1);
@@ -135,31 +137,31 @@
goalw Perm.thy [inj_def,id_def] "!!A B. A<=B ==> id(A): inj(A,B)";
by (REPEAT (ares_tac [CollectI,lam_type] 1));
by (etac subsetD 1 THEN assume_tac 1);
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
qed "id_subset_inj";
val id_inj = subset_refl RS id_subset_inj;
goalw Perm.thy [id_def,surj_def] "id(A): surj(A,A)";
-by (fast_tac (ZF_cs addIs [lam_type,beta]) 1);
+by (fast_tac (!claset addIs [lam_type,beta]) 1);
qed "id_surj";
goalw Perm.thy [bij_def] "id(A): bij(A,A)";
-by (fast_tac (ZF_cs addIs [id_inj,id_surj]) 1);
+by (fast_tac (!claset addIs [id_inj,id_surj]) 1);
qed "id_bij";
goalw Perm.thy [id_def] "A <= B <-> id(A) : A->B";
-by (fast_tac (ZF_cs addSIs [lam_type] addDs [apply_type] addss ZF_ss) 1);
+by (fast_tac (!claset addSIs [lam_type] addDs [apply_type] addss (!simpset)) 1);
qed "subset_iff_id";
(*** Converse of a function ***)
goalw Perm.thy [inj_def] "!!f. f: inj(A,B) ==> converse(f) : range(f)->A";
-by (asm_simp_tac (ZF_ss addsimps [Pi_iff, function_def]) 1);
+by (asm_simp_tac (!simpset addsimps [Pi_iff, function_def]) 1);
by (etac CollectE 1);
-by (asm_simp_tac (ZF_ss addsimps [apply_iff]) 1);
-by (fast_tac (ZF_cs addDs [fun_is_rel]) 1);
+by (asm_simp_tac (!simpset addsimps [apply_iff]) 1);
+by (fast_tac (!claset addDs [fun_is_rel]) 1);
qed "inj_converse_fun";
(** Equations for converse(f) **)
@@ -167,12 +169,12 @@
(*The premises are equivalent to saying that f is injective...*)
val prems = goal Perm.thy
"[| f: A->B; converse(f): C->A; a: A |] ==> converse(f)`(f`a) = a";
-by (fast_tac (ZF_cs addIs (prems@[apply_Pair,apply_equality,converseI])) 1);
+by (fast_tac (!claset addIs (prems@[apply_Pair,apply_equality,converseI])) 1);
qed "left_inverse_lemma";
goal Perm.thy
"!!f. [| f: inj(A,B); a: A |] ==> converse(f)`(f`a) = a";
-by (fast_tac (ZF_cs addIs [left_inverse_lemma,inj_converse_fun,inj_is_fun]) 1);
+by (fast_tac (!claset addIs [left_inverse_lemma,inj_converse_fun,inj_is_fun]) 1);
qed "left_inverse";
val left_inverse_bij = bij_is_inj RS left_inverse;
@@ -190,9 +192,12 @@
by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1));
qed "right_inverse";
+(*Cannot add [left_inverse, right_inverse] to default simpset: there are too
+ many ways of expressing sufficient conditions.*)
+
goal Perm.thy "!!f. [| f: bij(A,B); b: B |] ==> f`(converse(f)`b) = b";
-by (fast_tac (ZF_cs addss
- (ZF_ss addsimps [bij_def, right_inverse, surj_range])) 1);
+by (fast_tac (!claset addss
+ (!simpset addsimps [bij_def, right_inverse, surj_range])) 1);
qed "right_inverse_bij";
(** Converses of injections, surjections, bijections **)
@@ -211,9 +216,10 @@
qed "inj_converse_surj";
goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)";
-by (fast_tac (ZF_cs addEs [surj_range RS subst, inj_converse_inj,
+by (fast_tac (!claset addEs [surj_range RS subst, inj_converse_inj,
inj_converse_surj]) 1);
qed "bij_converse_bij";
+(*Adding this as an SI seems to cause looping*)
(** Composition of two relations **)
@@ -221,7 +227,7 @@
(*The inductive definition package could derive these theorems for (r O s)*)
goalw Perm.thy [comp_def] "!!r s. [| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "compI";
val prems = goalw Perm.thy [comp_def]
@@ -237,62 +243,64 @@
THEN prune_params_tac)
(read_instantiate [("xz","<a,c>")] compE);
-val comp_cs = ZF_cs addSIs [idI] addIs [compI] addSEs [compE,idE];
+AddSIs [idI];
+AddIs [compI];
+AddSEs [compE,idE];
(** Domain and Range -- see Suppes, section 3.1 **)
(*Boyer et al., Set Theory in First-Order Logic, JAR 2 (1986), 287-327*)
goal Perm.thy "range(r O s) <= range(r)";
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
qed "range_comp";
goal Perm.thy "!!r s. domain(r) <= range(s) ==> range(r O s) = range(r)";
by (rtac (range_comp RS equalityI) 1);
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
qed "range_comp_eq";
goal Perm.thy "domain(r O s) <= domain(s)";
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
qed "domain_comp";
goal Perm.thy "!!r s. range(s) <= domain(r) ==> domain(r O s) = domain(s)";
by (rtac (domain_comp RS equalityI) 1);
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
qed "domain_comp_eq";
goal Perm.thy "(r O s)``A = r``(s``A)";
-by (fast_tac (comp_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "image_comp";
(** Other results **)
goal Perm.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
qed "comp_mono";
(*composition preserves relations*)
goal Perm.thy "!!r s. [| s<=A*B; r<=B*C |] ==> (r O s) <= A*C";
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
qed "comp_rel";
(*associative law for composition*)
goal Perm.thy "(r O s) O t = r O (s O t)";
-by (fast_tac (comp_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "comp_assoc";
(*left identity of composition; provable inclusions are
id(A) O r <= r
and [| r<=A*B; B<=C |] ==> r <= id(C) O r *)
goal Perm.thy "!!r A B. r<=A*B ==> id(B) O r = r";
-by (fast_tac (comp_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "left_comp_id";
(*right identity of composition; provable inclusions are
r O id(A) <= r
and [| r<=A*B; A<=C |] ==> r <= r O id(C) *)
goal Perm.thy "!!r A B. r<=A*B ==> r O id(A) = r";
-by (fast_tac (comp_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "right_comp_id";
@@ -300,15 +308,15 @@
goalw Perm.thy [function_def]
"!!f g. [| function(g); function(f) |] ==> function(f O g)";
-by (fast_tac (ZF_cs addIs [compI] addSEs [compE, Pair_inject]) 1);
+by (fast_tac (!claset addIs [compI] addSEs [compE, Pair_inject]) 1);
qed "comp_function";
goal Perm.thy "!!f g. [| g: A->B; f: B->C |] ==> (f O g) : A->C";
by (asm_full_simp_tac
- (ZF_ss addsimps [Pi_def, comp_function, Pow_iff, comp_rel]
+ (!simpset addsimps [Pi_def, comp_function, Pow_iff, comp_rel]
setloop etac conjE) 1);
by (stac (range_rel_subset RS domain_comp_eq) 1 THEN assume_tac 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "comp_fun";
goal Perm.thy "!!f g. [| g: A->B; f: B->C; a:A |] ==> (f O g)`a = f`(g`a)";
@@ -316,6 +324,8 @@
apply_Pair,apply_type] 1));
qed "comp_fun_apply";
+Addsimps [comp_fun_apply];
+
(*Simplifies compositions of lambda-abstractions*)
val [prem] = goal Perm.thy
"[| !!x. x:A ==> b(x): B \
@@ -324,7 +334,7 @@
by (rtac comp_fun 1);
by (rtac lam_funtype 2);
by (typechk_tac (prem::ZF_typechecks));
-by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]
+by (asm_simp_tac (!simpset
setsolver type_auto_tac [lam_type, lam_funtype, prem]) 1);
qed "comp_lam";
@@ -332,18 +342,18 @@
by (res_inst_tac [("d", "%y. converse(g) ` (converse(f) ` y)")]
f_imp_injective 1);
by (REPEAT (ares_tac [comp_fun, inj_is_fun] 1));
-by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply, left_inverse]
+by (asm_simp_tac (!simpset addsimps [left_inverse]
setsolver type_auto_tac [inj_is_fun, apply_type]) 1);
qed "comp_inj";
goalw Perm.thy [surj_def]
"!!f g. [| g: surj(A,B); f: surj(B,C) |] ==> (f O g) : surj(A,C)";
-by (best_tac (ZF_cs addSIs [comp_fun,comp_fun_apply]) 1);
+by (best_tac (!claset addSIs [comp_fun,comp_fun_apply]) 1);
qed "comp_surj";
goalw Perm.thy [bij_def]
"!!f g. [| g: bij(A,B); f: bij(B,C) |] ==> (f O g) : bij(A,C)";
-by (fast_tac (ZF_cs addIs [comp_inj,comp_surj]) 1);
+by (fast_tac (!claset addIs [comp_inj,comp_surj]) 1);
qed "comp_bij";
@@ -353,26 +363,26 @@
goalw Perm.thy [inj_def]
"!!f g. [| (f O g): inj(A,C); g: A->B; f: B->C |] ==> g: inj(A,B)";
-by (safe_tac comp_cs);
+by (safe_tac (!claset));
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));
-by (asm_simp_tac (FOL_ss addsimps [comp_fun_apply]) 1);
+by (asm_simp_tac (!simpset ) 1);
qed "comp_mem_injD1";
goalw Perm.thy [inj_def,surj_def]
"!!f g. [| (f O g): inj(A,C); g: surj(A,B); f: B->C |] ==> f: inj(B,C)";
-by (safe_tac comp_cs);
+by (safe_tac (!claset));
by (res_inst_tac [("x1", "x")] (bspec RS bexE) 1);
by (eres_inst_tac [("x1", "w")] (bspec RS bexE) 3);
by (REPEAT (assume_tac 1));
-by (safe_tac comp_cs);
+by (safe_tac (!claset));
by (res_inst_tac [("t", "op `(g)")] subst_context 1);
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));
-by (asm_simp_tac (FOL_ss addsimps [comp_fun_apply]) 1);
+by (asm_simp_tac (!simpset ) 1);
qed "comp_mem_injD2";
goalw Perm.thy [surj_def]
"!!f g. [| (f O g): surj(A,C); g: A->B; f: B->C |] ==> f: surj(B,C)";
-by (fast_tac (comp_cs addSIs [comp_fun_apply RS sym, apply_type]) 1);
+by (best_tac (!claset addSIs [comp_fun_apply RS sym, apply_type]) 1);
qed "comp_mem_surjD1";
goal Perm.thy
@@ -382,10 +392,10 @@
goalw Perm.thy [inj_def,surj_def]
"!!f g. [| (f O g): surj(A,C); g: A->B; f: inj(B,C) |] ==> g: surj(A,B)";
-by (safe_tac comp_cs);
+by (safe_tac (!claset));
by (eres_inst_tac [("x1", "f`y")] (bspec RS bexE) 1);
by (REPEAT (ares_tac [apply_type] 1 ORELSE dtac comp_fun_applyD 1));
-by (best_tac (comp_cs addSIs [apply_type]) 1);
+by (best_tac (!claset addSIs [apply_type]) 1);
qed "comp_mem_surjD2";
@@ -394,9 +404,9 @@
(*left inverse of composition; one inclusion is
f: A->B ==> id(A) <= converse(f) O f *)
goalw Perm.thy [inj_def] "!!f. f: inj(A,B) ==> converse(f) O f = id(A)";
-by (fast_tac (comp_cs addIs [equalityI, apply_Pair]
+by (fast_tac (!claset addIs [equalityI, apply_Pair]
addEs [domain_type]
- addss (ZF_ss addsimps [apply_iff])) 1);
+ addss (!simpset addsimps [apply_iff])) 1);
qed "left_comp_inverse";
(*right inverse of composition; one inclusion is
@@ -407,8 +417,8 @@
val appfD = (prem RS CollectD1) RSN (3,apply_equality2);
by (cut_facts_tac [prem] 1);
by (rtac equalityI 1);
-by (best_tac (comp_cs addEs [domain_type, range_type, make_elim appfD]) 1);
-by (best_tac (comp_cs addIs [apply_Pair]) 1);
+by (best_tac (!claset addEs [domain_type, range_type, make_elim appfD]) 1);
+by (best_tac (!claset addIs [apply_Pair]) 1);
qed "right_comp_inverse";
(** Proving that a function is a bijection **)
@@ -416,18 +426,18 @@
goalw Perm.thy [id_def]
"!!f A B. [| f: A->B; g: B->A |] ==> \
\ f O g = id(B) <-> (ALL y:B. f`(g`y)=y)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (dres_inst_tac [("t", "%h.h`y ")] subst_context 1);
-by (asm_full_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1);
+by (Asm_full_simp_tac 1);
by (rtac fun_extension 1);
by (REPEAT (ares_tac [comp_fun, lam_type] 1));
-by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1);
+by (Auto_tac());
qed "comp_eq_id_iff";
goalw Perm.thy [bij_def]
"!!f A B. [| f: A->B; g: B->A; f O g = id(B); g O f = id(A) \
\ |] ==> f : bij(A,B)";
-by (asm_full_simp_tac (ZF_ss addsimps [comp_eq_id_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [comp_eq_id_iff]) 1);
by (REPEAT (ares_tac [conjI, f_imp_injective, f_imp_surjective] 1
ORELSE eresolve_tac [bspec, apply_type] 1));
qed "fg_imp_bijective";
@@ -437,7 +447,7 @@
qed "nilpotent_imp_bijective";
goal Perm.thy "!!f A B. [| converse(f): B->A; f: A->B |] ==> f : bij(A,B)";
-by (asm_simp_tac (ZF_ss addsimps [fg_imp_bijective, comp_eq_id_iff,
+by (asm_simp_tac (!simpset addsimps [fg_imp_bijective, comp_eq_id_iff,
left_inverse_lemma, right_inverse_lemma]) 1);
qed "invertible_imp_bijective";
@@ -450,9 +460,9 @@
by (res_inst_tac [("d","%z. if(z:B, converse(f)`z, converse(g)`z)")]
lam_injective 1);
by (ALLGOALS
- (asm_simp_tac (ZF_ss addsimps [inj_is_fun RS apply_type, left_inverse]
+ (asm_simp_tac (!simpset addsimps [inj_is_fun RS apply_type, left_inverse]
setloop (split_tac [expand_if] ORELSE' etac UnE))));
-by (fast_tac (ZF_cs addSEs [inj_is_fun RS apply_type] addDs [equals0D]) 1);
+by (fast_tac (!claset addSEs [inj_is_fun RS apply_type] addDs [equals0D]) 1);
qed "inj_disjoint_Un";
goalw Perm.thy [surj_def]
@@ -461,7 +471,7 @@
by (DEPTH_SOLVE_1 (eresolve_tac [fun_disjoint_apply1, fun_disjoint_apply2] 1
ORELSE ball_tac 1
ORELSE (rtac trans 1 THEN atac 2)
- ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1));
+ ORELSE step_tac (!claset addIs [fun_disjoint_Un]) 1));
qed "surj_disjoint_Un";
(*A simple, high-level proof; the version for injections follows from it,
@@ -480,7 +490,7 @@
val prems = goalw Perm.thy [surj_def]
"f: Pi(A,B) ==> f: surj(A, f``A)";
val rls = apply_equality :: (prems RL [apply_Pair,Pi_type]);
-by (fast_tac (ZF_cs addIs rls) 1);
+by (fast_tac (!claset addIs rls) 1);
qed "surj_image";
goal Perm.thy "!!f. [| f: Pi(C,B); A<=C |] ==> restrict(f,A)``A = f``A";
@@ -493,7 +503,7 @@
goalw Perm.thy [inj_def]
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): inj(C,B)";
-by (safe_tac (ZF_cs addSEs [restrict_type2]));
+by (safe_tac (!claset addSEs [restrict_type2]));
by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD,
box_equals, restrict] 1));
qed "restrict_inj";
@@ -507,7 +517,7 @@
goalw Perm.thy [inj_def,bij_def]
"!!f. [| f: inj(A,B); C<=A |] ==> restrict(f,C): bij(C, f``C)";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (REPEAT (eresolve_tac [bspec RS bspec RS mp, subsetD,
box_equals, restrict] 1
ORELSE ares_tac [surj_is_fun,restrict_surj] 1));
@@ -517,23 +527,24 @@
(*** Lemmas for Ramsey's Theorem ***)
goalw Perm.thy [inj_def] "!!f. [| f: inj(A,B); B<=D |] ==> f: inj(A,D)";
-by (fast_tac (ZF_cs addSEs [fun_weaken_type]) 1);
+by (fast_tac (!claset addSEs [fun_weaken_type]) 1);
qed "inj_weaken_type";
val [major] = goal Perm.thy
"[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})";
by (rtac (major RS restrict_bij RS bij_is_inj RS inj_weaken_type) 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (cut_facts_tac [major] 1);
by (rewtac inj_def);
-by (fast_tac (ZF_cs addEs [range_type, mem_irrefl] addDs [apply_equality]) 1);
+by (fast_tac (!claset addEs [range_type, mem_irrefl]
+ addDs [apply_equality]) 1);
qed "inj_succ_restrict";
goalw Perm.thy [inj_def]
"!!f. [| f: inj(A,B); a~:A; b~:B |] ==> \
\ cons(<a,b>,f) : inj(cons(a,A), cons(b,B))";
-by (fast_tac (ZF_cs addIs [apply_type]
- addss (ZF_ss addsimps [fun_extend, fun_extend_apply2,
+by (fast_tac (!claset addIs [apply_type]
+ addss (!simpset addsimps [fun_extend, fun_extend_apply2,
fun_extend_apply1]) ) 1);
qed "inj_extend";
--- a/src/ZF/Perm.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Perm.thy Fri Jan 03 15:01:55 1997 +0100
@@ -9,7 +9,7 @@
-- Lemmas for the Schroeder-Bernstein Theorem
*)
-Perm = ZF + "mono" +
+Perm = upair + mono + func +
consts
O :: [i,i]=>i (infixr 60)
--- a/src/ZF/QPair.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/QPair.ML Fri Jan 03 15:01:55 1997 +0100
@@ -24,12 +24,19 @@
(** Lemmas for showing that <a;b> uniquely determines a and b **)
+qed_goalw "QPair_empty" thy [QPair_def]
+ "<0;0> = 0"
+ (fn _=> [Simp_tac 1]);
+
qed_goalw "QPair_iff" thy [QPair_def]
"<a;b> = <c;d> <-> a=c & b=d"
(fn _=> [rtac sum_equal_iff 1]);
bind_thm ("QPair_inject", (QPair_iff RS iffD1 RS conjE));
+Addsimps [QPair_empty, QPair_iff];
+AddSEs [QPair_inject];
+
qed_goal "QPair_inject1" thy "<a;b> = <c;d> ==> a=c"
(fn [major]=>
[ (rtac (major RS QPair_inject) 1), (assume_tac 1) ]);
@@ -43,8 +50,10 @@
Generalizes Cartesian product ***)
qed_goalw "QSigmaI" thy [QSigma_def]
- "[| a:A; b:B(a) |] ==> <a;b> : QSigma(A,B)"
- (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
+ "!!A B. [| a:A; b:B(a) |] ==> <a;b> : QSigma(A,B)"
+ (fn _ => [ Fast_tac 1 ]);
+
+AddSIs [QSigmaI];
(*The general elimination rule*)
qed_goalw "QSigmaE" thy [QSigma_def]
@@ -70,44 +79,45 @@
(fn [major]=>
[ (rtac (major RS QSigmaE2) 1), (assume_tac 1) ]);
-val qpair_cs = ZF_cs addSIs [QSigmaI] addSEs [QSigmaE2, QSigmaE, QPair_inject];
+AddSEs [QSigmaE2, QSigmaE];
qed_goalw "QSigma_cong" thy [QSigma_def]
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \
\ QSigma(A,B) = QSigma(A',B')"
- (fn prems=> [ (simp_tac (ZF_ss addsimps prems) 1) ]);
+ (fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);
qed_goal "QSigma_empty1" thy "QSigma(0,B) = 0"
- (fn _ => [ (fast_tac (qpair_cs addIs [equalityI]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addIs [equalityI]) 1) ]);
qed_goal "QSigma_empty2" thy "A <*> 0 = 0"
- (fn _ => [ (fast_tac (qpair_cs addIs [equalityI]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addIs [equalityI]) 1) ]);
+
+Addsimps [QSigma_empty1, QSigma_empty2];
(*** Projections: qfst, qsnd ***)
qed_goalw "qfst_conv" thy [qfst_def] "qfst(<a;b>) = a"
(fn _=>
- [ (fast_tac (qpair_cs addIs [the_equality]) 1) ]);
+ [ (fast_tac (!claset addIs [the_equality]) 1) ]);
qed_goalw "qsnd_conv" thy [qsnd_def] "qsnd(<a;b>) = b"
(fn _=>
- [ (fast_tac (qpair_cs addIs [the_equality]) 1) ]);
+ [ (fast_tac (!claset addIs [the_equality]) 1) ]);
-val qpair_ss = ZF_ss addsimps [qfst_conv,qsnd_conv];
+Addsimps [qfst_conv, qsnd_conv];
qed_goal "qfst_type" thy
"!!p. p:QSigma(A,B) ==> qfst(p) : A"
- (fn _=> [ (fast_tac (qpair_cs addss qpair_ss) 1) ]);
+ (fn _=> [ Auto_tac() ]);
qed_goal "qsnd_type" thy
"!!p. p:QSigma(A,B) ==> qsnd(p) : B(qfst(p))"
- (fn _=> [ (fast_tac (qpair_cs addss qpair_ss) 1) ]);
+ (fn _=> [ Auto_tac() ]);
goal thy "!!a A B. a: QSigma(A,B) ==> <qfst(a); qsnd(a)> = a";
-by (etac QSigmaE 1);
-by (asm_simp_tac qpair_ss 1);
+by (Auto_tac());
qed "QPair_qfst_qsnd_eq";
@@ -116,30 +126,30 @@
(*A META-equality, so that it applies to higher types as well...*)
qed_goalw "qsplit" thy [qsplit_def]
"qsplit(%x y.c(x,y), <a;b>) == c(a,b)"
- (fn _ => [ (simp_tac qpair_ss 1),
+ (fn _ => [ (Simp_tac 1),
(rtac reflexive_thm 1) ]);
+Addsimps [qsplit];
+
qed_goal "qsplit_type" thy
"[| p:QSigma(A,B); \
\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x;y>) \
\ |] ==> qsplit(%x y.c(x,y), p) : C(p)"
(fn major::prems=>
[ (rtac (major RS QSigmaE) 1),
- (asm_simp_tac (qpair_ss addsimps (qsplit::prems)) 1) ]);
+ (asm_simp_tac (!simpset addsimps prems) 1) ]);
goalw thy [qsplit_def]
"!!u. u: A<*>B ==> \
\ R(qsplit(c,u)) <-> (ALL x:A. ALL y:B. u = <x;y> --> R(c(x,y)))";
-by (etac QSigmaE 1);
-by (asm_simp_tac qpair_ss 1);
-by (fast_tac qpair_cs 1);
+by (Auto_tac());
qed "expand_qsplit";
(*** qsplit for predicates: result type o ***)
goalw thy [qsplit_def] "!!R a b. R(a,b) ==> qsplit(R, <a;b>)";
-by (asm_simp_tac qpair_ss 1);
+by (Asm_simp_tac 1);
qed "qsplitI";
val major::sigma::prems = goalw thy [qsplit_def]
@@ -148,11 +158,12 @@
\ |] ==> P";
by (rtac (sigma RS QSigmaE) 1);
by (cut_facts_tac [major] 1);
-by (asm_full_simp_tac (qpair_ss addsimps prems) 1);
+by (REPEAT (ares_tac prems 1));
+by (Asm_full_simp_tac 1);
qed "qsplitE";
goalw thy [qsplit_def] "!!R a b. qsplit(R,<a;b>) ==> R(a,b)";
-by (asm_full_simp_tac qpair_ss 1);
+by (Asm_full_simp_tac 1);
qed "qsplitD";
@@ -160,11 +171,11 @@
qed_goalw "qconverseI" thy [qconverse_def]
"!!a b r. <a;b>:r ==> <b;a>:qconverse(r)"
- (fn _ => [ (fast_tac qpair_cs 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
qed_goalw "qconverseD" thy [qconverse_def]
"!!a b r. <a;b> : qconverse(r) ==> <b;a> : r"
- (fn _ => [ (fast_tac qpair_cs 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
qed_goalw "qconverseE" thy [qconverse_def]
"[| yx : qconverse(r); \
@@ -176,22 +187,22 @@
(hyp_subst_tac 1),
(assume_tac 1) ]);
-val qconverse_cs = qpair_cs addSIs [qconverseI]
- addSEs [qconverseD,qconverseE];
+AddSIs [qconverseI];
+AddSEs [qconverseD, qconverseE];
qed_goal "qconverse_qconverse" thy
"!!A B r. r<=QSigma(A,B) ==> qconverse(qconverse(r)) = r"
- (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addSIs [equalityI]) 1) ]);
qed_goal "qconverse_type" thy
"!!A B r. r <= A <*> B ==> qconverse(r) <= B <*> A"
- (fn _ => [ (fast_tac qconverse_cs 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
qed_goal "qconverse_prod" thy "qconverse(A <*> B) = B <*> A"
- (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addSIs [equalityI]) 1) ]);
qed_goal "qconverse_empty" thy "qconverse(0) = 0"
- (fn _ => [ (fast_tac (qconverse_cs addSIs [equalityI]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addSIs [equalityI]) 1) ]);
(**** The Quine-inspired notion of disjoint sum ****)
@@ -201,11 +212,11 @@
(** Introduction rules for the injections **)
goalw thy qsum_defs "!!a A B. a : A ==> QInl(a) : A <+> B";
-by (REPEAT (ares_tac [UnI1,QSigmaI,singletonI] 1));
+by (Fast_tac 1);
qed "QInlI";
goalw thy qsum_defs "!!b A B. b : B ==> QInr(b) : A <+> B";
-by (REPEAT (ares_tac [UnI2,QSigmaI,singletonI] 1));
+by (Fast_tac 1);
qed "QInrI";
(** Elimination rules **)
@@ -220,24 +231,30 @@
ORELSE eresolve_tac (prems@[QSigmaE,singletonE,ssubst]) 1));
qed "qsumE";
+AddSIs [QInlI, QInrI];
+
(** Injection and freeness equivalences, for rewriting **)
goalw thy qsum_defs "QInl(a)=QInl(b) <-> a=b";
-by (simp_tac (ZF_ss addsimps [QPair_iff]) 1);
+by (Simp_tac 1);
qed "QInl_iff";
goalw thy qsum_defs "QInr(a)=QInr(b) <-> a=b";
-by (simp_tac (ZF_ss addsimps [QPair_iff]) 1);
+by (Simp_tac 1);
qed "QInr_iff";
goalw thy qsum_defs "QInl(a)=QInr(b) <-> False";
-by (simp_tac (ZF_ss addsimps [QPair_iff, one_not_0 RS not_sym]) 1);
+by (Simp_tac 1);
qed "QInl_QInr_iff";
goalw thy qsum_defs "QInr(b)=QInl(a) <-> False";
-by (simp_tac (ZF_ss addsimps [QPair_iff, one_not_0]) 1);
+by (Simp_tac 1);
qed "QInr_QInl_iff";
+goalw thy qsum_defs "0<+>0 = 0";
+by (Simp_tac 1);
+qed "qsum_empty";
+
(*Injection and freeness rules*)
bind_thm ("QInl_inject", (QInl_iff RS iffD1));
@@ -245,45 +262,46 @@
bind_thm ("QInl_neq_QInr", (QInl_QInr_iff RS iffD1 RS FalseE));
bind_thm ("QInr_neq_QInl", (QInr_QInl_iff RS iffD1 RS FalseE));
-val qsum_cs =
- qpair_cs addSIs [PartI, QInlI, QInrI]
- addSEs [PartE, qsumE, QInl_neq_QInr, QInr_neq_QInl]
- addSDs [QInl_inject, QInr_inject];
+AddSEs [qsumE, QInl_neq_QInr, QInr_neq_QInl];
+AddSDs [QInl_inject, QInr_inject];
+Addsimps [QInl_iff, QInr_iff, QInl_QInr_iff, QInr_QInl_iff, qsum_empty];
goal thy "!!A B. QInl(a): A<+>B ==> a: A";
-by (fast_tac qsum_cs 1);
+by (Fast_tac 1);
qed "QInlD";
goal thy "!!A B. QInr(b): A<+>B ==> b: B";
-by (fast_tac qsum_cs 1);
+by (Fast_tac 1);
qed "QInrD";
(** <+> is itself injective... who cares?? **)
goal thy
"u: A <+> B <-> (EX x. x:A & u=QInl(x)) | (EX y. y:B & u=QInr(y))";
-by (fast_tac qsum_cs 1);
+by (Fast_tac 1);
qed "qsum_iff";
goal thy "A <+> B <= C <+> D <-> A<=C & B<=D";
-by (fast_tac qsum_cs 1);
+by (Fast_tac 1);
qed "qsum_subset_iff";
goal thy "A <+> B = C <+> D <-> A=C & B=D";
-by (simp_tac (ZF_ss addsimps [extension,qsum_subset_iff]) 1);
-by (fast_tac ZF_cs 1);
+by (simp_tac (!simpset addsimps [extension,qsum_subset_iff]) 1);
+by (Fast_tac 1);
qed "qsum_equal_iff";
(*** Eliminator -- qcase ***)
goalw thy qsum_defs "qcase(c, d, QInl(a)) = c(a)";
-by (simp_tac (ZF_ss addsimps [qsplit, cond_0]) 1);
+by (Simp_tac 1);
qed "qcase_QInl";
goalw thy qsum_defs "qcase(c, d, QInr(b)) = d(b)";
-by (simp_tac (ZF_ss addsimps [qsplit, cond_1]) 1);
+by (Simp_tac 1);
qed "qcase_QInr";
+Addsimps [qcase_QInl, qcase_QInr];
+
val major::prems = goal thy
"[| u: A <+> B; \
\ !!x. x: A ==> c(x): C(QInl(x)); \
@@ -291,24 +309,23 @@
\ |] ==> qcase(c,d,u) : C(u)";
by (rtac (major RS qsumE) 1);
by (ALLGOALS (etac ssubst));
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps
- (prems@[qcase_QInl,qcase_QInr]))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "qcase_type";
(** Rules for the Part primitive **)
goal thy "Part(A <+> B,QInl) = {QInl(x). x: A}";
-by (fast_tac (qsum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_QInl";
goal thy "Part(A <+> B,QInr) = {QInr(y). y: B}";
-by (fast_tac (qsum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_QInr";
goal thy "Part(A <+> B, %x.QInr(h(x))) = {QInr(y). y: Part(B,h)}";
-by (fast_tac (qsum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_QInr2";
goal thy "!!A B C. C <= A <+> B ==> Part(C,QInl) Un Part(C,QInr) = C";
-by (fast_tac (qsum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_qsum_equality";
--- a/src/ZF/QPair.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/QPair.thy Fri Jan 03 15:01:55 1997 +0100
@@ -11,7 +11,7 @@
1966.
*)
-QPair = Sum + "simpdata" +
+QPair = Sum +
consts
QPair :: [i, i] => i ("<(_;/ _)>")
qfst,qsnd :: i => i
--- a/src/ZF/QUniv.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/QUniv.ML Fri Jan 03 15:01:55 1997 +0100
@@ -20,7 +20,7 @@
val [prem] = goalw QUniv.thy [sum_def]
"Transset(C) ==> (A+B) Int C <= (A Int C) + (B Int C)";
by (stac Int_Un_distrib 1);
-by (fast_tac (ZF_cs addSDs [prem RS Transset_Pair_D]) 1);
+by (fast_tac (!claset addSDs [prem RS Transset_Pair_D]) 1);
qed "Transset_sum_Int_subset";
(** Introduction and elimination rules avoid tiresome folding/unfolding **)
@@ -170,7 +170,7 @@
\ a: Vfrom(X,i) & b: Vfrom(X,i)";
by (dtac (Transset_Vfrom_succ RS equalityD1 RS subsetD RS PowD) 1);
by (assume_tac 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "doubleton_in_Vfrom_D";
(*This weaker version says a, b exist at the same level*)
@@ -187,13 +187,13 @@
goalw Univ.thy [Pair_def]
"!!X. [| <a,b> : Vfrom(X,succ(i)); Transset(X) |] ==> \
\ a: Vfrom(X,i) & b: Vfrom(X,i)";
-by (fast_tac (ZF_cs addSDs [doubleton_in_Vfrom_D, Vfrom_doubleton_D]) 1);
+by (fast_tac (!claset addSDs [doubleton_in_Vfrom_D, Vfrom_doubleton_D]) 1);
qed "Pair_in_Vfrom_D";
goal Univ.thy
"!!X. Transset(X) ==> \
\ (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))";
-by (fast_tac (ZF_cs addSDs [Pair_in_Vfrom_D]) 1);
+by (fast_tac (!claset addSDs [Pair_in_Vfrom_D]) 1);
qed "product_Int_Vfrom_subset";
(*** Intersecting <a;b> with Vfrom... ***)
@@ -227,11 +227,11 @@
by (etac Ord_cases 1 THEN REPEAT_FIRST hyp_subst_tac);
(*0 case*)
by (stac Vfrom_0 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
(*succ(j) case*)
by (rtac (Transset_0 RS QPair_Int_Vfrom_succ_subset RS subset_trans) 1);
by (rtac (succI1 RS UN_upper) 1);
(*Limit(i) case*)
-by (asm_simp_tac (ZF_ss addsimps [Limit_Vfrom_eq, Int_UN_distrib, subset_refl,
+by (asm_simp_tac (!simpset addsimps [Limit_Vfrom_eq, Int_UN_distrib, subset_refl,
UN_mono, QPair_Int_Vset_subset_trans]) 1);
qed "QPair_Int_Vset_subset_UN";
--- a/src/ZF/QUniv.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/QUniv.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
A small universe for lazy recursive types
*)
-QUniv = Univ + QPair + "mono" + "equalities" +
+QUniv = Univ + QPair + mono + equalities +
consts
quniv :: i=>i
--- a/src/ZF/ROOT.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ROOT.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,6 +11,8 @@
val banner = "ZF Set Theory (in FOL)";
writeln banner;
+eta_contract:=false;
+
(*For Pure/tactic?? A crude way of adding structure to rules*)
fun CHECK_SOLVED tac st =
case Sequence.pull (tac st) of
@@ -32,6 +34,8 @@
use "thy_syntax.ML";
use_thy "Let";
+use_thy "func";
+use "typechk.ML";
use_thy "InfDatatype";
use_thy "List";
use_thy "EquivClass";
--- a/src/ZF/Rel.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Rel.ML Fri Jan 03 15:01:55 1997 +0100
@@ -30,7 +30,7 @@
qed "symI";
goalw Rel.thy [sym_def] "!!r. [| sym(r); <x,y>: r |] ==> <y,x>: r";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "symE";
(* antisymmetry *)
@@ -43,18 +43,18 @@
val prems = goalw Rel.thy [antisym_def]
"!!r. [| antisym(r); <x,y>: r; <y,x>: r |] ==> x=y";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "antisymE";
(* transitivity *)
goalw Rel.thy [trans_def]
"!!r. [| trans(r); <a,b>:r; <b,c>:r |] ==> <a,c>:r";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "transD";
goalw Rel.thy [trans_on_def]
"!!r. [| trans[A](r); <a,b>:r; <b,c>:r; a:A; b:A; c:A |] ==> <a,c>:r";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "trans_onD";
--- a/src/ZF/Rel.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Rel.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
Relations in Zermelo-Fraenkel Set Theory
*)
-Rel = ZF +
+Rel = domrange +
consts
refl,irrefl,equiv :: [i,i]=>o
sym,asym,antisym,trans :: i=>o
--- a/src/ZF/Resid/Confluence.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Confluence.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,11 +13,11 @@
goalw Confluence.thy [confluence_def]
"!!u.[|confluence(Spar_red)|]==> confluence(Sred)";
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
by (dres_inst_tac [("x4","x"),("x3","y"),("x1","z")]
(spec RS spec RS mp RS spec RS mp) 1);
-by (ALLGOALS(asm_simp_tac (reduc_ss addsimps [red_par_red])));
-by (fast_tac (reduc_cs addIs [par_red_red]) 1);
+by (ALLGOALS(asm_simp_tac (!simpset addsimps [red_par_red])));
+by (fast_tac (!claset addIs [par_red_red]) 1);
val lemma1 = result();
(* ------------------------------------------------------------------------- *)
@@ -28,8 +28,8 @@
"!!u.[|confluence(Spar_red1)|]==> strip";
by (resolve_tac [impI RS allI RS allI] 1);
by (etac Spar_red.induct 1);
-by (fast_tac reduc_cs 1);
-by (fast_tac (ZF_cs addIs [Spar_red.trans]) 1);
+by (Fast_tac 1);
+by (fast_tac (!claset addIs [Spar_red.trans]) 1);
val strip_lemma_r = result();
@@ -37,12 +37,12 @@
"!!u.strip==> confluence(Spar_red)";
by (resolve_tac [impI RS allI RS allI] 1);
by (etac Spar_red.induct 1);
-by (fast_tac reduc_cs 1);
-by (step_tac ZF_cs 1);
+by (Fast_tac 1);
+by (step_tac (!claset) 1);
by (dres_inst_tac [("x1","z")] (spec RS mp) 1);
by (REPEAT(eresolve_tac [exE,conjE] 2));
by (dres_inst_tac [("x1","ua")] (spec RS mp) 2);
-by (fast_tac (ZF_cs addIs [Spar_red.trans]) 3);
+by (fast_tac (!claset addIs [Spar_red.trans]) 3);
by (TRYALL assume_tac );
val strip_lemma_l = result();
@@ -59,15 +59,15 @@
goalw Confluence.thy [confluence_def] "confluence(Spar_red1)";
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
by (forward_tac [simulation] 1);
by (forw_inst_tac [("n","z")] simulation 1);
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
by (forw_inst_tac [("v","va")] paving 1);
by (TRYALL assume_tac);
-by (REPEAT(step_tac ZF_cs 1));
+by (REPEAT(step_tac (!claset) 1));
by (res_inst_tac [("v","vu")] completeness 1);
-by (ALLGOALS(asm_simp_tac (reduc_ss addsimps [completeness])));
+by (ALLGOALS(asm_simp_tac (!simpset addsimps [completeness])));
val parallel_moves = result();
goal Confluence.thy "confluence(Spar_red)";
--- a/src/ZF/Resid/Conversion.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Conversion.ML Fri Jan 03 15:01:55 1997 +0100
@@ -7,14 +7,14 @@
open Conversion;
-val conv_ss = reduc_ss addsimps (Sconv1.intrs@[Sconv.one_step,Sconv.refl]);
+val (!simpset) = (!simpset) addsimps (Sconv1.intrs@[Sconv.one_step,Sconv.refl]);
goal Conversion.thy
"!!u.m<--->n ==> n<--->m";
by (etac Sconv.induct 1);
by (etac Sconv1.induct 1);
by (resolve_tac [Sconv.trans] 4);
-by (ALLGOALS(asm_simp_tac (conv_ss) ));
+by (ALLGOALS(asm_simp_tac (!simpset) ));
val conv_sym = result();
(* ------------------------------------------------------------------------- *)
@@ -25,18 +25,18 @@
"!!u.m<--->n ==> EX p.(m --->p) & (n ---> p)";
by (etac Sconv.induct 1);
by (etac Sconv1.induct 1);
-by (fast_tac (reduc_cs addIs [red1D1,redD2]) 1);
-by (fast_tac (reduc_cs addIs [red1D1,redD2]) 1);
-by (fast_tac (reduc_cs addIs [red1D1,redD2]) 1);
+by (fast_tac (!claset addIs [red1D1,redD2]) 1);
+by (fast_tac (!claset addIs [red1D1,redD2]) 1);
+by (fast_tac (!claset addIs [red1D1,redD2]) 1);
by (cut_facts_tac [confluence_beta_reduction] 1);
by (rewtac confluence_def);
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
by (dres_inst_tac [("x4","n"),("x3","pa"),("x1","pb")]
(spec RS spec RS mp RS spec RS mp) 1);
by (TRYALL assume_tac);
-by (step_tac ZF_cs 1);
-by (step_tac ZF_cs 1);
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
+by (step_tac (!claset) 1);
+by (step_tac (!claset) 1);
by (res_inst_tac [("n","pa")]Sred.trans 1);
by (res_inst_tac [("n","pb"),("p","ua")] Sred.trans 3);
by (TRYALL assume_tac);
--- a/src/ZF/Resid/Cube.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Cube.ML Fri Jan 03 15:01:55 1997 +0100
@@ -22,16 +22,16 @@
\ regular(u)-->(ALL w.w~v-->w~u--> \
\ w |> u = (w|>v) |> (u|>v))";
by (etac Ssub.induct 1);
-by (asm_simp_tac (prism_ss) 1);
-by (asm_simp_tac (prism_ss ) 1);
-by (asm_simp_tac (prism_ss ) 1);
+by (asm_simp_tac prism_ss 1);
+by (asm_simp_tac prism_ss 1);
+by (asm_simp_tac prism_ss 1);
by (dresolve_tac [spec RS mp RS mp] 1 THEN resolve_tac [Scomp.Comp_Fun] 1
THEN resolve_tac [Scomp.Comp_Fun] 2 THEN (REPEAT(assume_tac 1)));
-by (asm_full_simp_tac (prism_ss ) 1);
-by (asm_simp_tac (prism_ss ) 1);
+by (asm_full_simp_tac prism_ss 1);
+by (asm_simp_tac prism_ss 1);
by (dresolve_tac [spec RS mp RS mp] 1 THEN resolve_tac [Scomp.Comp_Fun] 1
THEN resolve_tac [Scomp.Comp_Fun] 2 THEN (REPEAT(assume_tac 1)));
-by (asm_full_simp_tac (prism_ss ) 1);
+by (asm_full_simp_tac prism_ss 1);
val prism_l = result();
goal Cube.thy
@@ -40,7 +40,7 @@
by (resolve_tac [prism_l RS mp RS spec RS mp RS mp ] 1);
by (rtac comp_trans 4);
by (assume_tac 4);
-by (ALLGOALS(asm_simp_tac (prism_ss)));
+by (ALLGOALS(asm_simp_tac prism_ss));
val prism = result();
@@ -57,12 +57,12 @@
by (etac comp_sym 1);
by (atac 1);
by (stac (prism RS sym) 1);
-by (asm_full_simp_tac (res1_ss addsimps
- [prism RS sym,union_l,union_preserve_regular,
- comp_sym_iff, union_sym]) 4);
-by (asm_full_simp_tac (res1_ss addsimps [union_r, comp_sym_iff]) 1);
-by (asm_full_simp_tac (res1_ss addsimps
- [union_preserve_regular, comp_sym_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps
+ [prism RS sym,union_l,union_preserve_regular,
+ comp_sym_iff, union_sym]) 4);
+by (asm_full_simp_tac (!simpset addsimps [union_r, comp_sym_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps
+ [union_preserve_regular, comp_sym_iff]) 1);
by (etac comp_trans 1);
by (atac 1);
val cube = result();
@@ -77,7 +77,7 @@
\ EX uv.EX vu. (w|>u) |> vu = (w|>v) |> uv & (w|>u)~vu &\
\ regular(vu) & (w|>v)~uv & regular(uv) ";
by (forw_inst_tac [("x1","v")] (comp_sym RS comp_trans) 1);
-by (REPEAT(step_tac ZF_cs 1));
+by (REPEAT(Step_tac 1));
by (rtac cube 1);
by (asm_simp_tac (prism_ss addsimps [comp_sym_iff]) 1);
by (asm_simp_tac (prism_ss addsimps [comp_sym_iff]) 1);
--- a/src/ZF/Resid/Redex.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Redex.ML Fri Jan 03 15:01:55 1997 +0100
@@ -27,8 +27,7 @@
by (simp_tac (rank_ss addsimps redexes.con_defs) 1);
val redex_rec_App = result();
-val redexes_ss = (arith_ss addsimps
- ([redex_rec_App,redex_rec_Fun,redex_rec_Var]@redexes.intrs));
+Addsimps ([redex_rec_App,redex_rec_Fun,redex_rec_Var]@redexes.intrs);
--- a/src/ZF/Resid/Reduction.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Reduction.ML Fri Jan 03 15:01:55 1997 +0100
@@ -22,18 +22,16 @@
val par_redD2 = Spar_red.dom_subset RS subsetD RS SigmaD2;
-val reduc_cs = res_cs
- addIs (Sred1.intrs@[Sred.one_step,Sred.refl]@Spar_red1.intrs@
- [Spar_red.one_step,lambda.dom_subset RS subsetD,
- unmark_type]@lambda.intrs@bool_typechecks)
- addSEs [Spar_red1.mk_cases redexes.con_defs "Fun(t) =1=> Fun(u)"];
+AddIs (Sred1.intrs@[Sred.one_step, Sred.refl]@Spar_red1.intrs@
+ [Spar_red.one_step, lambda.dom_subset RS subsetD,
+ unmark_type]@lambda.intrs@bool_typechecks);
+AddSEs [Spar_red1.mk_cases redexes.con_defs "Fun(t) =1=> Fun(u)"];
-val reduc_ss = term_ss addsimps
- (Sred1.intrs@[Sred.one_step,Sred.refl]@Spar_red1.intrs@
- [Spar_red.one_step,substL_type,redD1,redD2,par_redD1,
- par_redD2,par_red1D2,unmark_type]);
+Addsimps (Sred1.intrs@[Sred.one_step, Sred.refl]@Spar_red1.intrs@
+ [Spar_red.one_step, substL_type, redD1, redD2, par_redD1,
+ par_redD2, par_red1D2, unmark_type]);
-val reducL_ss = reduc_ss setloop (SELECT_GOAL (safe_tac reduc_cs));
+val reducL_ss = !simpset setloop (SELECT_GOAL (safe_tac (!claset)));
(* ------------------------------------------------------------------------- *)
(* Lemmas for reduction *)
@@ -42,34 +40,34 @@
goal Reduction.thy "!!u. m--->n ==> Fun(m) ---> Fun(n)";
by (etac Sred.induct 1);
by (resolve_tac [Sred.trans] 3);
-by (ALLGOALS (asm_simp_tac reduc_ss ));
+by (ALLGOALS (Asm_simp_tac ));
val red_Fun = result();
goal Reduction.thy
"!!u.[|n:lambda; m ---> m'|]==> Apl(m,n)--->Apl(m',n)";
by (etac Sred.induct 1);
by (resolve_tac [Sred.trans] 3);
-by (ALLGOALS (asm_simp_tac reduc_ss ));
+by (ALLGOALS (Asm_simp_tac ));
val red_Apll = result();
goal Reduction.thy
"!!u.[|n:lambda; m ---> m'|]==> Apl(n,m)--->Apl(n,m')";
by (etac Sred.induct 1);
by (resolve_tac [Sred.trans] 3);
-by (ALLGOALS (asm_simp_tac reduc_ss ));
+by (ALLGOALS (Asm_simp_tac ));
val red_Aplr = result();
goal Reduction.thy
"!!u.[|m ---> m'; n--->n'|]==> Apl(m,n)--->Apl(m',n')";
by (res_inst_tac [("n","Apl(m',n)")] Sred.trans 1);
-by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [red_Apll,red_Aplr]) ));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [red_Apll,red_Aplr]) ));
val red_Apl = result();
goal Reduction.thy
"!!u.[|m:lambda; m':lambda; n:lambda; n':lambda; m ---> m'; n--->n'|]==> \
\ Apl(Fun(m),n)---> n'/m'";
by (res_inst_tac [("n","Apl(Fun(m'),n')")] Sred.trans 1);
-by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [red_Apl,red_Fun]) ));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [red_Apl,red_Fun]) ));
val red_beta = result();
@@ -80,25 +78,25 @@
goal Reduction.thy "!!u.m:lambda==> m =1=> m";
by (eresolve_tac [lambda.induct] 1);
-by (ALLGOALS (asm_simp_tac reduc_ss ));
+by (ALLGOALS (Asm_simp_tac ));
val refl_par_red1 = result();
goal Reduction.thy "!!u.m-1->n ==> m=1=>n";
by (etac Sred1.induct 1);
-by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [refl_par_red1]) ));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [refl_par_red1]) ));
val red1_par_red1 = result();
goal Reduction.thy "!!u.m--->n ==> m===>n";
by (etac Sred.induct 1);
by (resolve_tac [Spar_red.trans] 3);
-by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [refl_par_red1,red1_par_red1]) ));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [refl_par_red1,red1_par_red1]) ));
val red_par_red = result();
goal Reduction.thy "!!u.m===>n ==> m--->n";
by (etac Spar_red.induct 1);
by (etac Spar_red1.induct 1);
by (resolve_tac [Sred.trans] 5);
-by (ALLGOALS (asm_simp_tac (reduc_ss addsimps [red_Fun,red_beta,red_Apl]) ));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [red_Fun,red_beta,red_Apl]) ));
val par_red_red = result();
@@ -109,10 +107,10 @@
goal Reduction.thy
"!!u.m=1=>n ==> EX v. m|>v = n & m~v & regular(v)";
by (etac Spar_red1.induct 1);
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
by (ALLGOALS (REPEAT o (resolve_tac [exI,conjI])));
by (TRYALL(resolve_tac [res_redex,res_App,res_Fun,res_Var]));
-by (ALLGOALS (asm_simp_tac (reduc_ss)));
+by (ALLGOALS (asm_simp_tac (!simpset)));
val simulation = result();
@@ -124,7 +122,7 @@
"!!u.u:redexes ==> \
\ ALL k:nat. unmark(lift_rec(u,k)) = lift_rec(unmark(u),k)";
by (eresolve_tac [redexes.induct] 1);
-by (ALLGOALS (asm_full_simp_tac (addsplit reduc_ss)));
+by (ALLGOALS (asm_full_simp_tac (addsplit (!simpset))));
val unmmark_lift_rec = result();
goal Reduction.thy
@@ -132,7 +130,7 @@
\ unmark(subst_rec(u,v,k)) = subst_rec(unmark(u),unmark(v),k)";
by (eresolve_tac [redexes.induct] 1);
by (ALLGOALS (asm_full_simp_tac
- ((addsplit reduc_ss) addsimps [unmmark_lift_rec])));
+ ((addsplit (!simpset)) addsimps [unmmark_lift_rec])));
val unmmark_subst_rec = result();
@@ -151,5 +149,5 @@
goal Reduction.thy
"!!u.[|u:lambda; u~v; regular(v)|]==> u =1=> unmark(u|>v)";
by (dres_inst_tac [("u1","u")] (completeness_l RS mp) 1);
-by (ALLGOALS (asm_full_simp_tac (reduc_ss addsimps [lambda_unmark]) ));
+by (ALLGOALS (asm_full_simp_tac (!simpset addsimps [lambda_unmark]) ));
val completeness = result();
--- a/src/ZF/Resid/Residuals.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Residuals.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,30 +11,29 @@
(* Setting up rule lists *)
(* ------------------------------------------------------------------------- *)
-val res_cs = union_cs
- addIs (Sres.intrs@redexes.intrs@Sreg.intrs@
- [subst_type]@nat_typechecks)
- addSEs (redexes.free_SEs @
- (map (Sres.mk_cases redexes.con_defs)
- ["residuals(Var(n),Var(n),v)",
- "residuals(Fun(t),Fun(u),v)",
- "residuals(App(b, u1, u2), App(0, v1, v2),v)",
- "residuals(App(b, u1, u2), App(1, Fun(v1), v2),v)",
- "residuals(Var(n),u,v)",
- "residuals(Fun(t),u,v)",
- "residuals(App(b, u1, u2), w,v)",
- "residuals(u,Var(n),v)",
- "residuals(u,Fun(t),v)",
- "residuals(w,App(b, u1, u2),v)"]) @
- (map (Ssub.mk_cases redexes.con_defs)
- ["Var(n) <== u",
- "Fun(n) <== u",
- "u <== Fun(n)",
- "App(1,Fun(t),a) <== u",
- "App(0,t,a) <== u"]) @
- [redexes.mk_cases redexes.con_defs "Fun(t):redexes"]);
+AddIs (Sres.intrs@redexes.intrs@Sreg.intrs@
+ [subst_type]@nat_typechecks);
+AddSEs (redexes.free_SEs @
+ (map (Sres.mk_cases redexes.con_defs)
+ ["residuals(Var(n),Var(n),v)",
+ "residuals(Fun(t),Fun(u),v)",
+ "residuals(App(b, u1, u2), App(0, v1, v2),v)",
+ "residuals(App(b, u1, u2), App(1, Fun(v1), v2),v)",
+ "residuals(Var(n),u,v)",
+ "residuals(Fun(t),u,v)",
+ "residuals(App(b, u1, u2), w,v)",
+ "residuals(u,Var(n),v)",
+ "residuals(u,Fun(t),v)",
+ "residuals(w,App(b, u1, u2),v)"]) @
+ (map (Ssub.mk_cases redexes.con_defs)
+ ["Var(n) <== u",
+ "Fun(n) <== u",
+ "u <== Fun(n)",
+ "App(1,Fun(t),a) <== u",
+ "App(0,t,a) <== u"]) @
+ [redexes.mk_cases redexes.con_defs "Fun(t):redexes"]);
-val res_ss = subst_ss addsimps (Sres.intrs);
+Addsimps Sres.intrs;
val resD1 =Sres.dom_subset RS subsetD RS SigmaD1;
val resD2 =Sres.dom_subset RS subsetD RS SigmaD2 RS SigmaD1;
val resD3 =Sres.dom_subset RS subsetD RS SigmaD2 RS SigmaD2;
@@ -47,14 +46,14 @@
goal Residuals.thy
"!!u.residuals(u,v,w) ==> ALL w1.residuals(u,v,w1) --> w1 = w";
by (etac Sres.induct 1);
-by (ALLGOALS (fast_tac res_cs ));
+by (ALLGOALS Fast_tac);
val residuals_function = result();
val res_is_func = residuals_function RS spec RS mp;
goal Residuals.thy
"!!u.u~v ==> regular(v) --> (EX w.residuals(u,v,w))";
by (etac Scomp.induct 1);
-by (ALLGOALS (fast_tac res_cs));
+by (ALLGOALS Fast_tac);
val residuals_intro = result();
val prems = goal Residuals.thy
@@ -64,7 +63,7 @@
by (resolve_tac prems 1);
by (resolve_tac [residuals_intro RS mp RS exE] 1);
by (resolve_tac [the_equality RS ssubst] 3);
-by (ALLGOALS (fast_tac (res_cs addIs [res_is_func])));
+by (ALLGOALS (fast_tac (!claset addIs [res_is_func])));
val comp_resfuncE = result();
@@ -72,8 +71,8 @@
(* Residual function *)
(* ------------------------------------------------------------------------- *)
-val resfunc_cs = (res_cs addIs [the_equality,res_is_func]
- addSEs [comp_resfuncE]);
+val resfunc_cs = (!claset addIs [the_equality,res_is_func]
+ addSEs [comp_resfuncE]);
goalw Residuals.thy [res_func_def]
"!!n.n:nat ==> Var(n) |> Var(n) = Var(n)";
@@ -101,20 +100,20 @@
"!!n.[|s~t; regular(t)|]==> regular(t) --> s |> t:redexes";
by (etac Scomp.induct 1);
by (ALLGOALS (asm_full_simp_tac
- (res_ss addsimps [res_Var,res_Fun,res_App,res_redex]
- setloop (SELECT_GOAL (safe_tac res_cs)))));
+ (!simpset addsimps [res_Var,res_Fun,res_App,res_redex]
+ setloop (SELECT_GOAL (safe_tac (!claset))))));
by (dres_inst_tac [("psi", "Fun(?u) |> ?v : redexes")] asm_rl 1);
by (asm_full_simp_tac
- (res_ss addsimps [res_Fun]
- setloop (SELECT_GOAL (safe_tac res_cs))) 1);
+ (!simpset addsimps [res_Fun]
+ setloop (SELECT_GOAL (safe_tac (!claset)))) 1);
val resfunc_type = result();
-val res1_ss = (res_ss addsimps
- [res_Var,res_Fun,res_App,res_redex,lift_rec_preserve_comp,
- lift_rec_preserve_reg,subst_rec_preserve_comp,resfunc_type,
- subst_rec_preserve_reg]@redexes.free_iffs);
+Addsimps ([res_Var, res_Fun, res_App, res_redex, lift_rec_preserve_comp,
+ lift_rec_preserve_reg, subst_rec_preserve_comp, resfunc_type,
+ subst_rec_preserve_reg] @
+ redexes.free_iffs);
-val res1L_ss = res1_ss setloop (SELECT_GOAL (safe_tac res_cs));
+val res1L_ss = !simpset setloop (SELECT_GOAL (safe_tac (!claset)));
(* ------------------------------------------------------------------------- *)
(* Commutation theorem *)
@@ -123,7 +122,7 @@
goal Residuals.thy
"!!u.u<==v ==> u~v";
by (etac Ssub.induct 1);
-by (ALLGOALS (asm_simp_tac res1_ss));
+by (ALLGOALS Asm_simp_tac);
val sub_comp = result();
goal Residuals.thy
@@ -136,10 +135,10 @@
"!!u.[|u~v; k:nat|]==> regular(v)--> (ALL n:nat. \
\ lift_rec(u,n) |> lift_rec(v,n) = lift_rec(u |> v,n))";
by (etac Scomp.induct 1);
-by (step_tac res_cs 1);
-by (ALLGOALS (asm_full_simp_tac ((addsplit res1_ss) addsimps [lift_subst])));
+by (Step_tac 1);
+by (ALLGOALS (asm_full_simp_tac ((addsplit (!simpset)) addsimps [lift_subst])));
by (dres_inst_tac [("psi", "ALL x:nat.Fun(?u(x)) |> ?v(x) = ?w(x)")] asm_rl 1);
-by (asm_full_simp_tac res1_ss 1);
+by (Asm_full_simp_tac 1);
val residuals_lift_rec = result();
goal Residuals.thy
@@ -147,18 +146,18 @@
\ (ALL n:nat.subst_rec(v1,u1,n) |> subst_rec(v2,u2,n) = \
\ subst_rec(v1 |> v2, u1 |> u2,n))";
by (etac Scomp.induct 1);
-by (step_tac (res_cs) 1);
+by (Step_tac 1);
by (ALLGOALS
- (asm_full_simp_tac ((addsplit res1_ss) addsimps ([residuals_lift_rec]))));
+ (asm_full_simp_tac ((addsplit (!simpset)) addsimps ([residuals_lift_rec]))));
by (dres_inst_tac [("psi", "ALL x.?P(x)")] asm_rl 1);
-by (asm_full_simp_tac (res1_ss addsimps ([substitution])) 1);
+by (asm_full_simp_tac (!simpset addsimps ([substitution])) 1);
val residuals_subst_rec = result();
goal Residuals.thy
"!!u.[|u1~u2; v1~v2; regular(u2); regular(v2)|]==>\
\ (v1/u1) |> (v2/u2) = (v1 |> v2)/(u1 |> u2)";
-by (asm_simp_tac (res1_ss addsimps ([residuals_subst_rec])) 1);
+by (asm_simp_tac (!simpset addsimps ([residuals_subst_rec])) 1);
val commutation = result();
(* ------------------------------------------------------------------------- *)
@@ -168,7 +167,7 @@
goal Residuals.thy
"!!u.u~v ==> ALL w. u~w --> v~w --> regular(w) --> (u|>w) ~ (v|>w)";
by (etac Scomp.induct 1);
-by (ALLGOALS (asm_simp_tac (res1L_ss )));
+by (ALLGOALS (asm_simp_tac res1L_ss));
by (dresolve_tac [spec RS mp RS mp RS mp] 1
THEN resolve_tac Scomp.intrs 1 THEN resolve_tac Scomp.intrs 2
THEN resolve_tac Sreg.intrs 3);
@@ -180,7 +179,7 @@
goal Residuals.thy
"!!u.u~v ==> regular(u) --> regular(v) --> regular(u|>v)";
by (etac Scomp.induct 1);
-by (safe_tac res_cs);
+by (safe_tac (!claset));
by (TRYALL (dres_inst_tac [("psi", "regular(Fun(?u) |> ?v)")] asm_rl));
by (ALLGOALS (asm_full_simp_tac res1L_ss));
val residuals_preserve_reg = result();
@@ -192,17 +191,17 @@
goal Residuals.thy
"!!u.u~v ==> v ~ (u un v)";
by (etac Scomp.induct 1);
-by (ALLGOALS (asm_full_simp_tac res_ss));
+by (ALLGOALS Asm_full_simp_tac);
val union_preserve_comp = result();
goal Residuals.thy
"!!u.u ~ v ==> regular(v) --> u|>v = (u un v)|>v";
by (etac Scomp.induct 1);
-by (safe_tac res_cs);
+by (safe_tac (!claset));
by (dres_inst_tac [("psi", "Fun(?u) |> ?v = ?w")] asm_rl 3);
by (ALLGOALS (asm_full_simp_tac (res1L_ss addsimps
[union_preserve_comp,comp_sym_iff])));
-by (asm_full_simp_tac (res1_ss addsimps
+by (asm_full_simp_tac (!simpset addsimps
[union_preserve_comp RS comp_sym,
comp_sym RS union_preserve_comp RS comp_sym]) 1);
val preservation1 = result();
--- a/src/ZF/Resid/SubUnion.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/SubUnion.ML Fri Jan 03 15:01:55 1997 +0100
@@ -23,42 +23,41 @@
goalw SubUnion.thy [union_def]
"!!u.n:nat==>Var(n) un Var(n)=Var(n)";
-by (asm_simp_tac redexes_ss 1);
+by (Asm_simp_tac 1);
by (simp_tac (rank_ss addsimps redexes.con_defs) 1);
val union_Var = result();
goalw SubUnion.thy [union_def]
"!!u.[|u:redexes; v:redexes|]==>Fun(u) un Fun(v)=Fun(u un v)";
-by (asm_simp_tac redexes_ss 1);
+by (Asm_simp_tac 1);
by (simp_tac (rank_ss addsimps redexes.con_defs) 1);
val union_Fun = result();
goalw SubUnion.thy [union_def]
"!!u.[|b1:bool; b2:bool; u1:redexes; v1:redexes; u2:redexes; v2:redexes|]==> \
\ App(b1,u1,v1) un App(b2,u2,v2)=App(b1 or b2,u1 un u2,v1 un v2)";
-by (asm_simp_tac redexes_ss 1);
+by (Asm_simp_tac 1);
by (simp_tac (rank_ss addsimps redexes.con_defs) 1);
val union_App = result();
-val union_ss = redexes_ss addsimps
- (Ssub.intrs@bool_simps@bool_typechecks@
- Sreg.intrs@Scomp.intrs@
- [or_1 RSN (3,or_commute RS trans),
- or_0 RSN (3,or_commute RS trans),
- union_App,union_Fun,union_Var,compD2,compD1,regD]);
+Addsimps (Ssub.intrs@bool_typechecks@
+ Sreg.intrs@Scomp.intrs@
+ [or_1 RSN (3,or_commute RS trans),
+ or_0 RSN (3,or_commute RS trans),
+ union_App,union_Fun,union_Var,compD2,compD1,regD]);
-val union_cs = (ZF_cs addIs Scomp.intrs addSEs
- [Sreg.mk_cases redexes.con_defs "regular(App(b,f,a))",
- Sreg.mk_cases redexes.con_defs "regular(Fun(b))",
- Sreg.mk_cases redexes.con_defs "regular(Var(b))",
- Scomp.mk_cases redexes.con_defs "Fun(u) ~ Fun(t)",
- Scomp.mk_cases redexes.con_defs "u ~ Fun(t)",
- Scomp.mk_cases redexes.con_defs "u ~ Var(n)",
- Scomp.mk_cases redexes.con_defs "u ~ App(b,t,a)",
- Scomp.mk_cases redexes.con_defs "Fun(t) ~ v",
- Scomp.mk_cases redexes.con_defs "App(b,f,a) ~ v",
- Scomp.mk_cases redexes.con_defs "Var(n) ~ u"
- ]);
+AddIs Scomp.intrs;
+AddSEs [Sreg.mk_cases redexes.con_defs "regular(App(b,f,a))",
+ Sreg.mk_cases redexes.con_defs "regular(Fun(b))",
+ Sreg.mk_cases redexes.con_defs "regular(Var(b))",
+ Scomp.mk_cases redexes.con_defs "Fun(u) ~ Fun(t)",
+ Scomp.mk_cases redexes.con_defs "u ~ Fun(t)",
+ Scomp.mk_cases redexes.con_defs "u ~ Var(n)",
+ Scomp.mk_cases redexes.con_defs "u ~ App(b,t,a)",
+ Scomp.mk_cases redexes.con_defs "Fun(t) ~ v",
+ Scomp.mk_cases redexes.con_defs "App(b,f,a) ~ v",
+ Scomp.mk_cases redexes.con_defs "Var(n) ~ u"
+ ];
@@ -68,25 +67,25 @@
goal SubUnion.thy "!!u.u:redexes ==> u ~ u";
by (eresolve_tac [redexes.induct] 1);
-by (ALLGOALS(fast_tac union_cs));
+by (ALLGOALS Fast_tac);
val comp_refl = result();
goal SubUnion.thy
"!!u.u ~ v ==> v ~ u";
by (etac Scomp.induct 1);
-by (ALLGOALS(fast_tac union_cs));
+by (ALLGOALS Fast_tac);
val comp_sym = result();
goal SubUnion.thy
"u ~ v <-> v ~ u";
-by (fast_tac (ZF_cs addIs [comp_sym]) 1);
+by (fast_tac (!claset addIs [comp_sym]) 1);
val comp_sym_iff = result();
goal SubUnion.thy
"!!u.u ~ v ==> ALL w.v ~ w-->u ~ w";
by (etac Scomp.induct 1);
-by (ALLGOALS(fast_tac union_cs));
+by (ALLGOALS Fast_tac);
val comp_trans1 = result();
val comp_trans = comp_trans1 RS spec RS mp;
@@ -99,20 +98,20 @@
"!!u.u ~ v ==> u <== (u un v)";
by (etac Scomp.induct 1);
by (etac boolE 3);
-by (ALLGOALS(asm_full_simp_tac union_ss));
+by (ALLGOALS Asm_full_simp_tac);
val union_l = result();
goal SubUnion.thy
"!!u.u ~ v ==> v <== (u un v)";
by (etac Scomp.induct 1);
by (eres_inst_tac [("c","b2")] boolE 3);
-by (ALLGOALS(asm_full_simp_tac union_ss));
+by (ALLGOALS Asm_full_simp_tac);
val union_r = result();
goal SubUnion.thy
"!!u.u ~ v ==> u un v = v un u";
by (etac Scomp.induct 1);
-by (ALLGOALS(asm_simp_tac (union_ss addsimps [or_commute])));
+by (ALLGOALS(asm_simp_tac (!simpset addsimps [or_commute])));
val union_sym = result();
(* ------------------------------------------------------------------------- *)
@@ -123,7 +122,7 @@
"!!u.u ~ v ==> regular(u)-->regular(v)-->regular(u un v)";
by (etac Scomp.induct 1);
by (ALLGOALS(asm_full_simp_tac
- (union_ss setloop(SELECT_GOAL (safe_tac union_cs)))));
+ (!simpset setloop(SELECT_GOAL (safe_tac (!claset))))));
by (dres_inst_tac [("psi", "regular(Fun(?u) un ?v)")] asm_rl 1);
-by (asm_full_simp_tac union_ss 1);
+by (Asm_full_simp_tac 1);
val union_preserve_regular = result();
--- a/src/ZF/Resid/Substitution.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Substitution.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,32 +13,30 @@
goal Arith.thy
"!!m.[| p < n; n:nat|]==> n~=p";
-by (fast_tac lt_cs 1);
+by (Fast_tac 1);
val gt_not_eq = result();
val succ_pred = prove_goal Arith.thy
"!!i.[|j:nat; i:nat|]==> i < j --> succ(j #- 1) = j"
(fn prems =>[(etac nat_induct 1),
- (fast_tac lt_cs 1),
- (asm_simp_tac arith_ss 1)]);
+ (Fast_tac 1),
+ (Asm_simp_tac 1)]);
goal Arith.thy
"!!i.[|succ(x)<n; n:nat; x:nat|]==> x < n#-1 ";
by (rtac succ_leE 1);
by (forward_tac [nat_into_Ord RS le_refl RS lt_trans] 1 THEN assume_tac 1);
-by (asm_simp_tac (arith_ss addsimps [succ_pred]) 1);
+by (asm_simp_tac (!simpset addsimps [succ_pred]) 1);
val lt_pred = result();
goal Arith.thy
"!!i.[|n < succ(x); p<n; p:nat; n:nat; x:nat|]==> n#-1 < x ";
by (rtac succ_leE 1);
-by (asm_simp_tac (arith_ss addsimps [succ_pred]) 1);
+by (asm_simp_tac (!simpset addsimps [succ_pred]) 1);
val gt_pred = result();
-val lift_ss = (union_ss addsimps
- [add_0_right,add_succ_right,nat_into_Ord,
- not_lt_iff_le,if_P,if_not_P]);
+Addsimps [nat_into_Ord, not_lt_iff_le, if_P, if_not_P];
(* ------------------------------------------------------------------------- *)
@@ -46,29 +44,29 @@
(* ------------------------------------------------------------------------- *)
goalw Substitution.thy [lift_rec_def]
"!!n.[|n:nat; i:nat|]==> lift_rec(Var(i),n) =if(i<n,Var(i),Var(succ(i)))";
-by (asm_full_simp_tac lift_ss 1);
+by (Asm_full_simp_tac 1);
val lift_rec_Var = result();
goalw Substitution.thy [lift_rec_def]
"!!n.[|n:nat; i:nat; k:nat; k le i|]==> lift_rec(Var(i),k) = Var(succ(i))";
-by (asm_full_simp_tac lift_ss 1);
+by (Asm_full_simp_tac 1);
val lift_rec_le = result();
goalw Substitution.thy [lift_rec_def]
"!!n.[|i:nat; k:nat; i<k |]==> lift_rec(Var(i),k) = Var(i)";
-by (asm_full_simp_tac lift_ss 1);
+by (Asm_full_simp_tac 1);
val lift_rec_gt = result();
goalw Substitution.thy [lift_rec_def]
"!!n.[|n:nat; k:nat|]==> \
\ lift_rec(Fun(t),k) = Fun(lift_rec(t,succ(k)))";
-by (asm_full_simp_tac lift_ss 1);
+by (Asm_full_simp_tac 1);
val lift_rec_Fun = result();
goalw Substitution.thy [lift_rec_def]
"!!n.[|n:nat; k:nat|]==> \
\ lift_rec(App(b,f,a),k) = App(b,lift_rec(f,k),lift_rec(a,k))";
-by (asm_full_simp_tac lift_ss 1);
+by (Asm_full_simp_tac 1);
val lift_rec_App = result();
(* ------------------------------------------------------------------------- *)
@@ -78,36 +76,36 @@
goalw Substitution.thy [subst_rec_def]
"!!n.[|i:nat; k:nat; u:redexes|]==> \
\ subst_rec(u,Var(i),k) = if(k<i,Var(i#-1),if(k=i,u,Var(i)))";
-by (asm_full_simp_tac (lift_ss addsimps [gt_not_eq,leI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [gt_not_eq,leI]) 1);
val subst_Var = result();
goalw Substitution.thy [subst_rec_def]
"!!n.[|n:nat; u:redexes|]==> subst_rec(u,Var(n),n) = u";
-by (asm_full_simp_tac (lift_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
val subst_eq = result();
goalw Substitution.thy [subst_rec_def]
"!!n.[|n:nat; u:redexes; p:nat; p<n|]==> \
\ subst_rec(u,Var(n),p) = Var(n#-1)";
-by (asm_full_simp_tac (lift_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
val subst_gt = result();
goalw Substitution.thy [subst_rec_def]
"!!n.[|n:nat; u:redexes; p:nat; n<p|]==> \
\ subst_rec(u,Var(n),p) = Var(n)";
-by (asm_full_simp_tac (lift_ss addsimps [gt_not_eq,leI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [gt_not_eq,leI]) 1);
val subst_lt = result();
goalw Substitution.thy [subst_rec_def]
"!!n.[|p:nat; u:redexes|]==> \
\ subst_rec(u,Fun(t),p) = Fun(subst_rec(lift(u),t,succ(p))) ";
-by (asm_full_simp_tac (lift_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
val subst_Fun = result();
goalw Substitution.thy [subst_rec_def]
"!!n.[|p:nat; u:redexes|]==> \
\ subst_rec(u,App(b,f,a),p) = App(b,subst_rec(u,f,p),subst_rec(u,a,p))";
-by (asm_full_simp_tac (lift_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
val subst_App = result();
fun addsplit ss = (ss setloop (split_inside_tac [expand_if])
@@ -118,7 +116,7 @@
"!!n.u:redexes ==> ALL k:nat.lift_rec(u,k):redexes";
by (eresolve_tac [redexes.induct] 1);
by (ALLGOALS(asm_full_simp_tac
- ((addsplit lift_ss) addsimps [lift_rec_Fun,lift_rec_App])));
+ ((addsplit (!simpset)) addsimps [lift_rec_Fun,lift_rec_App])));
val lift_rec_type_a = result();
val lift_rec_type = lift_rec_type_a RS bspec;
@@ -126,16 +124,15 @@
"!!n.v:redexes ==> ALL n:nat.ALL u:redexes.subst_rec(u,v,n):redexes";
by (eresolve_tac [redexes.induct] 1);
by (ALLGOALS(asm_full_simp_tac
- ((addsplit lift_ss) addsimps [subst_Fun,subst_App,
+ ((addsplit (!simpset)) addsimps [subst_Fun,subst_App,
lift_rec_type])));
val subst_type_a = result();
val subst_type = subst_type_a RS bspec RS bspec;
-val subst_ss = (lift_ss addsimps
- [subst_eq,subst_gt,subst_lt,subst_Fun,subst_App,subst_type,
- lift_rec_le,lift_rec_gt,lift_rec_Fun,lift_rec_App,
- lift_rec_type]);
+Addsimps [subst_eq, subst_gt, subst_lt, subst_Fun, subst_App, subst_type,
+ lift_rec_le, lift_rec_gt, lift_rec_Fun, lift_rec_App,
+ lift_rec_type];
(* ------------------------------------------------------------------------- *)
@@ -146,19 +143,19 @@
"!!n.u:redexes ==> ALL n:nat.ALL i:nat.i le n --> \
\ (lift_rec(lift_rec(u,i),succ(n)) = lift_rec(lift_rec(u,n),i))";
by ((eresolve_tac [redexes.induct] 1)
- THEN (ALLGOALS(asm_simp_tac subst_ss)));
-by (step_tac ZF_cs 1);
+ THEN (ALLGOALS Asm_simp_tac));
+by (step_tac (!claset) 1);
by (excluded_middle_tac "na < xa" 1);
by ((forward_tac [lt_trans2] 2) THEN (assume_tac 2));
by (ALLGOALS(asm_full_simp_tac
- ((addsplit subst_ss) addsimps [leI])));
+ ((addsplit (!simpset)) addsimps [leI])));
val lift_lift_rec = result();
goalw Substitution.thy []
"!!n.[|u:redexes; n:nat|]==> \
\ lift_rec(lift(u),succ(n)) = lift(lift_rec(u,n))";
-by (asm_simp_tac (lift_ss addsimps [lift_lift_rec]) 1);
+by (asm_simp_tac (!simpset addsimps [lift_lift_rec]) 1);
val lift_lift = result();
goal Substitution.thy
@@ -167,24 +164,24 @@
\ lift_rec(subst_rec(u,v,n),m) = \
\ subst_rec(lift_rec(u,m),lift_rec(v,succ(m)),n)";
by ((eresolve_tac [redexes.induct] 1)
- THEN (ALLGOALS(asm_simp_tac (subst_ss addsimps [lift_lift]))));
-by (step_tac ZF_cs 1);
+ THEN (ALLGOALS(asm_simp_tac (!simpset addsimps [lift_lift]))));
+by (step_tac (!claset) 1);
by (excluded_middle_tac "na < x" 1);
-by (asm_full_simp_tac (subst_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
by (eres_inst_tac [("j","na")] leE 1);
-by (asm_full_simp_tac ((addsplit subst_ss)
+by (asm_full_simp_tac ((addsplit (!simpset))
addsimps [leI,gt_pred,succ_pred]) 1);
by (hyp_subst_tac 1);
-by (asm_full_simp_tac (subst_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
by (forw_inst_tac [("j","x")] lt_trans2 1);
by (assume_tac 1);
-by (asm_full_simp_tac (subst_ss addsimps [leI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [leI]) 1);
val lift_rec_subst_rec = result();
goalw Substitution.thy []
"!!n.[|v:redexes; u:redexes; n:nat|]==> \
\ lift_rec(u/v,n) = lift_rec(u,n)/lift_rec(v,succ(n))";
-by (asm_full_simp_tac (subst_ss addsimps [lift_rec_subst_rec]) 1);
+by (asm_full_simp_tac (!simpset addsimps [lift_rec_subst_rec]) 1);
val lift_subst = result();
@@ -194,19 +191,19 @@
\ lift_rec(subst_rec(u,v,n),m) = \
\ subst_rec(lift_rec(u,m),lift_rec(v,m),succ(n))";
by ((eresolve_tac [redexes.induct] 1)
- THEN (ALLGOALS(asm_simp_tac (subst_ss addsimps [lift_lift]))));
-by (step_tac ZF_cs 1);
+ THEN (ALLGOALS(asm_simp_tac (!simpset addsimps [lift_lift]))));
+by (step_tac (!claset) 1);
by (excluded_middle_tac "na < x" 1);
-by (asm_full_simp_tac (subst_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
by (eres_inst_tac [("i","x")] leE 1);
by (forward_tac [lt_trans1] 1 THEN assume_tac 1);
by (ALLGOALS(asm_full_simp_tac
- (subst_ss addsimps [succ_pred,leI,gt_pred])));
+ (!simpset addsimps [succ_pred,leI,gt_pred])));
by (hyp_subst_tac 1);
-by (asm_full_simp_tac (subst_ss addsimps [leI]) 1);
+by (asm_full_simp_tac (!simpset addsimps [leI]) 1);
by (excluded_middle_tac "na < xa" 1);
-by (asm_full_simp_tac (subst_ss) 1);
-by (asm_full_simp_tac (subst_ss addsimps [leI]) 1);
+by (asm_full_simp_tac (!simpset) 1);
+by (asm_full_simp_tac (!simpset addsimps [leI]) 1);
val lift_rec_subst_rec_lt = result();
@@ -214,12 +211,12 @@
"!!n.u:redexes ==> \
\ ALL n:nat.ALL v:redexes.subst_rec(v,lift_rec(u,n),n) = u";
by ((eresolve_tac [redexes.induct] 1)
- THEN (ALLGOALS(asm_simp_tac subst_ss)));
-by (step_tac ZF_cs 1);
+ THEN (ALLGOALS Asm_simp_tac));
+by (step_tac (!claset) 1);
by (excluded_middle_tac "na < x" 1);
(* x <= na *)
-by (asm_full_simp_tac (subst_ss) 1);
-by (asm_full_simp_tac (subst_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
+by (asm_full_simp_tac (!simpset) 1);
val subst_rec_lift_rec = result();
goal Substitution.thy
@@ -228,35 +225,35 @@
\ subst_rec(subst_rec(w,u,n),subst_rec(lift_rec(w,m),v,succ(n)),m)=\
\ subst_rec(w,subst_rec(u,v,m),n)";
by ((eresolve_tac [redexes.induct] 1) THEN
- (ALLGOALS(asm_simp_tac (subst_ss addsimps
+ (ALLGOALS(asm_simp_tac (!simpset addsimps
[lift_lift RS sym,lift_rec_subst_rec_lt]))));
-by (step_tac ZF_cs 1);
+by (step_tac (!claset) 1);
by (excluded_middle_tac "na le succ(xa)" 1);
-by (asm_full_simp_tac (subst_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
by (forward_tac [nat_into_Ord RS le_refl RS lt_trans] 1 THEN assume_tac 1);
by (etac leE 1);
-by (asm_simp_tac (subst_ss addsimps [succ_pred,lt_pred]) 2);
+by (asm_simp_tac (!simpset addsimps [succ_pred,lt_pred]) 2);
by (forward_tac [succ_leI RS lt_trans] 1 THEN assume_tac 1);
by (forw_inst_tac [("i","x")]
(nat_into_Ord RS le_refl RS lt_trans) 1 THEN assume_tac 1);
-by (asm_simp_tac (subst_ss addsimps [succ_pred,lt_pred]) 1);
+by (asm_simp_tac (!simpset addsimps [succ_pred,lt_pred]) 1);
by (eres_inst_tac [("i","na")] leE 1);
by (asm_full_simp_tac
- (subst_ss addsimps [succ_pred,subst_rec_lift_rec,leI]) 2);
+ (!simpset addsimps [succ_pred,subst_rec_lift_rec,leI]) 2);
by (excluded_middle_tac "na < x" 1);
-by (asm_full_simp_tac (subst_ss) 1);
+by (asm_full_simp_tac (!simpset) 1);
by (eres_inst_tac [("j","na")] leE 1);
-by (asm_simp_tac (subst_ss addsimps [gt_pred]) 1);
-by (asm_simp_tac (subst_ss addsimps [subst_rec_lift_rec]) 1);
+by (asm_simp_tac (!simpset addsimps [gt_pred]) 1);
+by (asm_simp_tac (!simpset addsimps [subst_rec_lift_rec]) 1);
by (forward_tac [lt_trans2] 1 THEN assume_tac 1);
-by (asm_simp_tac (subst_ss addsimps [gt_pred]) 1);
+by (asm_simp_tac (!simpset addsimps [gt_pred]) 1);
val subst_rec_subst_rec = result();
goalw Substitution.thy []
"!!n.[|v:redexes; u:redexes; w:redexes; n:nat|]==> \
\ subst_rec(w,u,n)/subst_rec(lift(w),v,succ(n)) = subst_rec(w,u/v,n)";
-by (asm_simp_tac (subst_ss addsimps [subst_rec_subst_rec]) 1);
+by (asm_simp_tac (!simpset addsimps [subst_rec_subst_rec]) 1);
val substitution = result();
(* ------------------------------------------------------------------------- *)
@@ -268,27 +265,27 @@
goal Substitution.thy
"!!n.[|n:nat; u ~ v|]==> ALL m:nat.lift_rec(u,m) ~ lift_rec(v,m)";
by (etac Scomp.induct 1);
-by (ALLGOALS(asm_simp_tac (subst_ss addsimps [comp_refl])));
+by (ALLGOALS(asm_simp_tac (!simpset addsimps [comp_refl])));
val lift_rec_preserve_comp = result();
goal Substitution.thy
"!!n.u2 ~ v2 ==> ALL m:nat.ALL u1:redexes.ALL v1:redexes.\
\ u1 ~ v1--> subst_rec(u1,u2,m) ~ subst_rec(v1,v2,m)";
by (etac Scomp.induct 1);
-by (ALLGOALS(asm_full_simp_tac ((addsplit subst_ss) addsimps
+by (ALLGOALS(asm_full_simp_tac ((addsplit (!simpset)) addsimps
([lift_rec_preserve_comp,comp_refl]))));
val subst_rec_preserve_comp = result();
goal Substitution.thy
"!!n.regular(u) ==> ALL m:nat.regular(lift_rec(u,m))";
by (eresolve_tac [Sreg.induct] 1);
-by (ALLGOALS(asm_full_simp_tac (addsplit subst_ss)));
+by (ALLGOALS(asm_full_simp_tac (addsplit (!simpset))));
val lift_rec_preserve_reg = result();
goal Substitution.thy
"!!n.regular(v) ==> \
\ ALL m:nat.ALL u:redexes.regular(u)-->regular(subst_rec(u,v,m))";
by (eresolve_tac [Sreg.induct] 1);
-by (ALLGOALS(asm_full_simp_tac ((addsplit subst_ss) addsimps
+by (ALLGOALS(asm_full_simp_tac ((addsplit (!simpset)) addsimps
[lift_rec_preserve_reg])));
val subst_rec_preserve_reg = result();
--- a/src/ZF/Resid/Terms.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Resid/Terms.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,17 +13,17 @@
goalw Terms.thy [unmark_def]
"unmark(Var(n)) = Var(n)";
-by (asm_simp_tac lift_ss 1);
+by (Asm_simp_tac 1);
val unmark_Var = result();
goalw Terms.thy [unmark_def]
"unmark(Fun(t)) = Fun(unmark(t))";
-by (asm_simp_tac lift_ss 1);
+by (Asm_simp_tac 1);
val unmark_Fun = result();
goalw Terms.thy [unmark_def]
"unmark(App(b,n,m)) = App(0,unmark(n),unmark(m))";
-by (asm_simp_tac lift_ss 1);
+by (Asm_simp_tac 1);
val unmark_App = result();
@@ -32,8 +32,9 @@
(* ------------------------------------------------------------------------- *)
-val term_ss = res1_ss addsimps ([unmark_App,unmark_Fun,unmark_Var,
- lambda.dom_subset RS subsetD]@lambda.intrs);
+Addsimps ([unmark_App, unmark_Fun, unmark_Var,
+ lambda.dom_subset RS subsetD] @
+ lambda.intrs);
(* ------------------------------------------------------------------------- *)
@@ -43,13 +44,13 @@
goalw Terms.thy [unmark_def]
"!!u.u:redexes ==> unmark(u):lambda";
by (eresolve_tac [redexes.induct] 1);
-by (ALLGOALS(asm_simp_tac term_ss));
+by (ALLGOALS Asm_simp_tac);
val unmark_type = result();
goal Terms.thy
"!!u.u:lambda ==> unmark(u) = u";
by (eresolve_tac [lambda.induct] 1);
-by (ALLGOALS(asm_simp_tac term_ss));
+by (ALLGOALS Asm_simp_tac);
val lambda_unmark = result();
@@ -60,14 +61,14 @@
goal Terms.thy
"!!u.[|v:lambda|]==> ALL k:nat.lift_rec(v,k):lambda";
by (eresolve_tac [lambda.induct] 1);
-by (ALLGOALS(asm_full_simp_tac (addsplit term_ss)));
+by (ALLGOALS(asm_full_simp_tac (addsplit (!simpset))));
val liftL_typea = result();
val liftL_type =liftL_typea RS bspec ;
goal Terms.thy
"!!n.[|v:lambda|]==> ALL n:nat.ALL u:lambda.subst_rec(u,v,n):lambda";
by (eresolve_tac [lambda.induct] 1);
-by (ALLGOALS(asm_full_simp_tac ((addsplit term_ss) addsimps [liftL_type])));
+by (ALLGOALS(asm_full_simp_tac ((addsplit (!simpset)) addsimps [liftL_type])));
val substL_typea = result();
val substL_type = substL_typea RS bspec RS bspec ;
--- a/src/ZF/Sum.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Sum.ML Fri Jan 03 15:01:55 1997 +0100
@@ -17,7 +17,7 @@
goalw Sum.thy [Part_def]
"!!a b A h. [| a : A; a=h(b) |] ==> a : Part(A,h)";
-by (REPEAT (ares_tac [exI,CollectI] 1));
+by (Fast_tac 1);
qed "Part_eqI";
val PartI = refl RSN (2,Part_eqI);
@@ -30,6 +30,9 @@
by (REPEAT (ares_tac prems 1));
qed "PartE";
+AddSIs [PartI];
+AddSEs [PartE];
+
goalw Sum.thy [Part_def] "Part(A,h) <= A";
by (rtac Collect_subset 1);
qed "Part_subset";
@@ -46,11 +49,11 @@
(** Introduction rules for the injections **)
goalw Sum.thy sum_defs "!!a A B. a : A ==> Inl(a) : A+B";
-by (REPEAT (ares_tac [UnI1,SigmaI,singletonI] 1));
+by (Fast_tac 1);
qed "InlI";
goalw Sum.thy sum_defs "!!b A B. b : B ==> Inr(b) : A+B";
-by (REPEAT (ares_tac [UnI2,SigmaI,singletonI] 1));
+by (Fast_tac 1);
qed "InrI";
(** Elimination rules **)
@@ -68,21 +71,25 @@
(** Injection and freeness equivalences, for rewriting **)
goalw Sum.thy sum_defs "Inl(a)=Inl(b) <-> a=b";
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
qed "Inl_iff";
goalw Sum.thy sum_defs "Inr(a)=Inr(b) <-> a=b";
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
qed "Inr_iff";
goalw Sum.thy sum_defs "Inl(a)=Inr(b) <-> False";
-by (simp_tac (ZF_ss addsimps [one_not_0 RS not_sym]) 1);
+by (Simp_tac 1);
qed "Inl_Inr_iff";
goalw Sum.thy sum_defs "Inr(b)=Inl(a) <-> False";
-by (simp_tac (ZF_ss addsimps [one_not_0]) 1);
+by (Simp_tac 1);
qed "Inr_Inl_iff";
+goalw Sum.thy sum_defs "0+0 = 0";
+by (Simp_tac 1);
+qed "sum_empty";
+
(*Injection and freeness rules*)
bind_thm ("Inl_inject", (Inl_iff RS iffD1));
@@ -90,29 +97,33 @@
bind_thm ("Inl_neq_Inr", (Inl_Inr_iff RS iffD1 RS FalseE));
bind_thm ("Inr_neq_Inl", (Inr_Inl_iff RS iffD1 RS FalseE));
-val sum_cs = ZF_cs addSIs [PartI, InlI, InrI]
- addSEs [PartE, sumE, Inl_neq_Inr, Inr_neq_Inl]
- addSDs [Inl_inject, Inr_inject];
+AddSIs [InlI, InrI];
+AddSEs [sumE, Inl_neq_Inr, Inr_neq_Inl];
+AddSDs [Inl_inject, Inr_inject];
+
+Addsimps [InlI, InrI, Inl_iff, Inr_iff, Inl_Inr_iff, Inr_Inl_iff, sum_empty];
+
+val sum_cs = !claset;
goal Sum.thy "!!A B. Inl(a): A+B ==> a: A";
-by (fast_tac sum_cs 1);
+by (Fast_tac 1);
qed "InlD";
goal Sum.thy "!!A B. Inr(b): A+B ==> b: B";
-by (fast_tac sum_cs 1);
+by (Fast_tac 1);
qed "InrD";
goal Sum.thy "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))";
-by (fast_tac sum_cs 1);
+by (Fast_tac 1);
qed "sum_iff";
goal Sum.thy "A+B <= C+D <-> A<=C & B<=D";
-by (fast_tac sum_cs 1);
+by (Fast_tac 1);
qed "sum_subset_iff";
goal Sum.thy "A+B = C+D <-> A=C & B=D";
-by (simp_tac (ZF_ss addsimps [extension,sum_subset_iff]) 1);
-by (fast_tac ZF_cs 1);
+by (simp_tac (!simpset addsimps [extension,sum_subset_iff]) 1);
+by (Fast_tac 1);
qed "sum_equal_iff";
goalw Sum.thy [sum_def] "A+A = 2*A";
@@ -123,13 +134,17 @@
(*** Eliminator -- case ***)
goalw Sum.thy sum_defs "case(c, d, Inl(a)) = c(a)";
-by (simp_tac (ZF_ss addsimps [cond_0]) 1);
+by (simp_tac (!simpset addsimps [cond_0]) 1);
qed "case_Inl";
goalw Sum.thy sum_defs "case(c, d, Inr(b)) = d(b)";
-by (simp_tac (ZF_ss addsimps [cond_1]) 1);
+by (simp_tac (!simpset addsimps [cond_1]) 1);
qed "case_Inr";
+Addsimps [case_Inl, case_Inr];
+
+val sum_ss = !simpset;
+
val major::prems = goal Sum.thy
"[| u: A+B; \
\ !!x. x: A ==> c(x): C(Inl(x)); \
@@ -137,8 +152,7 @@
\ |] ==> case(c,d,u) : C(u)";
by (rtac (major RS sumE) 1);
by (ALLGOALS (etac ssubst));
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps
- (prems@[case_Inl,case_Inr]))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "case_type";
goal Sum.thy
@@ -146,11 +160,7 @@
\ R(case(c,d,u)) <-> \
\ ((ALL x:A. u = Inl(x) --> R(c(x))) & \
\ (ALL y:B. u = Inr(y) --> R(d(y))))";
-by (etac sumE 1);
-by (asm_simp_tac (ZF_ss addsimps [case_Inl]) 1);
-by (fast_tac sum_cs 1);
-by (asm_simp_tac (ZF_ss addsimps [case_Inr]) 1);
-by (fast_tac sum_cs 1);
+by (Auto_tac());
qed "expand_case";
val major::prems = goal Sum.thy
@@ -159,39 +169,35 @@
\ !!y. y:B ==> d(y)=d'(y) \
\ |] ==> case(c,d,z) = case(c',d',z)";
by (resolve_tac [major RS sumE] 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps ([case_Inl, case_Inr] @ prems))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "case_cong";
-val [major] = goal Sum.thy
- "z: A+B ==> \
-\ case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = \
-\ case(%x. c(c'(x)), %y. d(d'(y)), z)";
-by (resolve_tac [major RS sumE] 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [case_Inl, case_Inr])));
+goal Sum.thy
+ "!!z. z: A+B ==> \
+\ case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = \
+\ case(%x. c(c'(x)), %y. d(d'(y)), z)";
+by (Auto_tac());
qed "case_case";
-val sum_ss = ZF_ss addsimps [InlI, InrI, Inl_iff, Inr_iff,
- Inl_Inr_iff, Inr_Inl_iff,
- case_Inl, case_Inr];
(*** More rules for Part(A,h) ***)
goal Sum.thy "!!A B h. A<=B ==> Part(A,h)<=Part(B,h)";
-by (fast_tac sum_cs 1);
+by (Fast_tac 1);
qed "Part_mono";
goal Sum.thy "Part(Collect(A,P), h) = Collect(Part(A,h), P)";
-by (fast_tac (sum_cs addSIs [equalityI]) 1);
+by (fast_tac (!claset addSIs [equalityI]) 1);
qed "Part_Collect";
bind_thm ("Part_CollectE", Part_Collect RS equalityD1 RS subsetD RS CollectE);
goal Sum.thy "Part(A+B,Inl) = {Inl(x). x: A}";
-by (fast_tac (sum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_Inl";
goal Sum.thy "Part(A+B,Inr) = {Inr(y). y: B}";
-by (fast_tac (sum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_Inr";
goalw Sum.thy [Part_def] "!!a. a : Part(A,h) ==> a : A";
@@ -199,13 +205,13 @@
qed "PartD1";
goal Sum.thy "Part(A,%x.x) = A";
-by (fast_tac (sum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_id";
goal Sum.thy "Part(A+B, %x.Inr(h(x))) = {Inr(y). y: Part(B,h)}";
-by (fast_tac (sum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_Inr2";
goal Sum.thy "!!A B C. C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C";
-by (fast_tac (sum_cs addIs [equalityI]) 1);
+by (fast_tac (!claset addIs [equalityI]) 1);
qed "Part_sum_equality";
--- a/src/ZF/Sum.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Sum.thy Fri Jan 03 15:01:55 1997 +0100
@@ -7,7 +7,7 @@
"Part" primitive for simultaneous recursive type definitions
*)
-Sum = Bool + "simpdata" +
+Sum = Bool + pair +
consts
"+" :: [i,i]=>i (infixr 65)
Inl,Inr :: i=>i
--- a/src/ZF/Trancl.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Trancl.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,7 +11,7 @@
goal Trancl.thy "bnd_mono(field(r)*field(r), %s. id(field(r)) Un (r O s))";
by (rtac bnd_monoI 1);
by (REPEAT (ares_tac [subset_refl, Un_mono, comp_mono] 2));
-by (fast_tac comp_cs 1);
+by (Fast_tac 1);
qed "rtrancl_bnd_mono";
val [prem] = goalw Trancl.thy [rtrancl_def] "r<=s ==> r^* <= s^*";
@@ -51,7 +51,7 @@
(*The premise ensures that r consists entirely of pairs*)
val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^*";
by (cut_facts_tac prems 1);
-by (fast_tac (ZF_cs addIs [r_into_rtrancl]) 1);
+by (fast_tac (!claset addIs [r_into_rtrancl]) 1);
qed "r_subset_rtrancl";
goal Trancl.thy "field(r^*) = field(r)";
@@ -68,7 +68,7 @@
\ !!x y z.[| P(<x,y>); <x,y>: r^*; <y,z>: r |] ==> P(<x,z>) |] \
\ ==> P(<a,b>)";
by (rtac ([rtrancl_def, rtrancl_bnd_mono, major] MRS def_induct) 1);
-by (fast_tac (ZF_cs addIs prems addSEs [idE,compE]) 1);
+by (fast_tac (!claset addIs prems) 1);
qed "rtrancl_full_induct";
(*nice induction rule.
@@ -85,7 +85,7 @@
by (EVERY1 [etac (spec RS mp), rtac refl]);
(*now do the induction*)
by (resolve_tac [major RS rtrancl_full_induct] 1);
-by (ALLGOALS (fast_tac (ZF_cs addIs prems)));
+by (ALLGOALS (fast_tac (!claset addIs prems)));
qed "rtrancl_induct";
(*transitivity of transitive closure!! -- by induction.*)
@@ -103,7 +103,7 @@
by (subgoal_tac "a = b | (EX y. <a,y> : r^* & <y,b> : r)" 1);
(*see HOL/trancl*)
by (rtac (major RS rtrancl_induct) 2);
-by (ALLGOALS (fast_tac (ZF_cs addSEs prems)));
+by (ALLGOALS (fast_tac (!claset addSEs prems)));
qed "rtranclE";
@@ -111,7 +111,7 @@
(*Transitivity of r^+ is proved by transitivity of r^* *)
goalw Trancl.thy [trans_def,trancl_def] "trans(r^+)";
-by (safe_tac comp_cs);
+by (safe_tac (!claset));
by (rtac (rtrancl_into_rtrancl RS (trans_rtrancl RS transD RS compI)) 1);
by (REPEAT (assume_tac 1));
qed "trans_trancl";
@@ -131,7 +131,7 @@
(*The premise ensures that r consists entirely of pairs*)
val prems = goal Trancl.thy "r <= Sigma(A,B) ==> r <= r^+";
by (cut_facts_tac prems 1);
-by (fast_tac (ZF_cs addIs [r_into_trancl]) 1);
+by (fast_tac (!claset addIs [r_into_trancl]) 1);
qed "r_subset_trancl";
(*intro rule by definition: from r^* and r *)
@@ -159,9 +159,9 @@
(*by induction on this formula*)
by (subgoal_tac "ALL z. <y,z> : r --> P(z)" 1);
(*now solve first subgoal: this formula is sufficient*)
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (etac rtrancl_induct 1);
-by (ALLGOALS (fast_tac (ZF_cs addIs (rtrancl_into_trancl1::prems))));
+by (ALLGOALS (fast_tac (!claset addIs (rtrancl_into_trancl1::prems))));
qed "trancl_induct";
(*elimination of r^+ -- NOT an induction rule*)
@@ -171,14 +171,14 @@
\ !!y.[| <a,y> : r^+; <y,b> : r |] ==> P \
\ |] ==> P";
by (subgoal_tac "<a,b> : r | (EX y. <a,y> : r^+ & <y,b> : r)" 1);
-by (fast_tac (ZF_cs addIs prems) 1);
+by (fast_tac (!claset addIs prems) 1);
by (rtac (rewrite_rule [trancl_def] major RS compEpair) 1);
by (etac rtranclE 1);
-by (ALLGOALS (fast_tac (ZF_cs addIs [rtrancl_into_trancl1])));
+by (ALLGOALS (fast_tac (!claset addIs [rtrancl_into_trancl1])));
qed "tranclE";
goalw Trancl.thy [trancl_def] "r^+ <= field(r)*field(r)";
-by (fast_tac (ZF_cs addEs [compE, rtrancl_type RS subsetD RS SigmaE2]) 1);
+by (fast_tac (!claset addEs [rtrancl_type RS subsetD RS SigmaE2]) 1);
qed "trancl_type";
val [prem] = goalw Trancl.thy [trancl_def] "r<=s ==> r^+ <= s^+";
--- a/src/ZF/Trancl.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Trancl.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
Transitive closure of a relation
*)
-Trancl = Fixedpt + Perm + "mono" + Rel +
+Trancl = Fixedpt + Perm + mono + Rel +
consts
rtrancl :: i=>i ("(_^*)" [100] 100) (*refl/transitive closure*)
trancl :: i=>i ("(_^+)" [100] 100) (*transitive closure*)
--- a/src/ZF/Univ.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Univ.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,7 +11,7 @@
(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*)
goal Univ.thy "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))";
by (stac (Vfrom_def RS def_transrec) 1);
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
qed "Vfrom";
(** Monotonicity **)
@@ -40,7 +40,7 @@
by (eps_ind_tac "x" 1);
by (stac Vfrom 1);
by (stac Vfrom 1);
-by (fast_tac (ZF_cs addSIs [rank_lt RS ltD]) 1);
+by (fast_tac (!claset addSIs [rank_lt RS ltD]) 1);
qed "Vfrom_rank_subset1";
goal Univ.thy "Vfrom(A,rank(x)) <= Vfrom(A,x)";
@@ -71,13 +71,13 @@
goal Univ.thy "!!x y. y:x ==> 0 : Vfrom(A,x)";
by (stac Vfrom 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "zero_in_Vfrom";
goal Univ.thy "i <= Vfrom(A,i)";
by (eps_ind_tac "i" 1);
by (stac Vfrom 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "i_subset_Vfrom";
goal Univ.thy "A <= Vfrom(A,i)";
@@ -89,20 +89,20 @@
goal Univ.thy "!!A a i. a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))";
by (stac Vfrom 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "subset_mem_Vfrom";
(** Finite sets and ordered pairs **)
goal Univ.thy "!!a. a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))";
by (rtac subset_mem_Vfrom 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
qed "singleton_in_Vfrom";
goal Univ.thy
"!!A. [| a: Vfrom(A,i); b: Vfrom(A,i) |] ==> {a,b} : Vfrom(A,succ(i))";
by (rtac subset_mem_Vfrom 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
qed "doubleton_in_Vfrom";
goalw Univ.thy [Pair_def]
@@ -160,7 +160,7 @@
(*opposite inclusion*)
by (rtac UN_least 1);
by (stac Vfrom 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Vfrom_Union";
(*** Vfrom applied to Limit ordinals ***)
@@ -270,7 +270,7 @@
goal Univ.thy "!!i A. Transset(A) ==> Transset(Vfrom(A,i))";
by (eps_ind_tac "i" 1);
by (stac Vfrom 1);
-by (fast_tac (ZF_cs addSIs [Transset_Union_family, Transset_Un,
+by (fast_tac (!claset addSIs [Transset_Union_family, Transset_Un,
Transset_Pow]) 1);
qed "Transset_Vfrom";
@@ -284,7 +284,7 @@
goalw Ordinal.thy [Pair_def,Transset_def]
"!!C. [| <a,b> <= C; Transset(C) |] ==> a: C & b: C";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Transset_Pair_subset";
goal Univ.thy
@@ -326,7 +326,7 @@
by (dtac Transset_Vfrom 1);
by (rtac subset_mem_Vfrom 1);
by (rewtac Transset_def);
-by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1);
+by (fast_tac (!claset addIs [Pair_in_Vfrom]) 1);
qed "prod_in_Vfrom";
val [aprem,bprem,limiti,transset] = goal Univ.thy
@@ -345,7 +345,7 @@
by (dtac Transset_Vfrom 1);
by (rtac subset_mem_Vfrom 1);
by (rewtac Transset_def);
-by (fast_tac (ZF_cs addIs [zero_in_Vfrom, Pair_in_Vfrom,
+by (fast_tac (!claset addIs [zero_in_Vfrom, Pair_in_Vfrom,
i_subset_Vfrom RS subsetD]) 1);
qed "sum_in_Vfrom";
@@ -371,7 +371,7 @@
by (rtac (succI1 RS UN_upper) 2);
by (rtac Pow_mono 1);
by (rewtac Transset_def);
-by (fast_tac (ZF_cs addIs [Pair_in_Vfrom]) 1);
+by (fast_tac (!claset addIs [Pair_in_Vfrom]) 1);
qed "fun_in_Vfrom";
val [aprem,bprem,limiti,transset] = goal Univ.thy
@@ -399,10 +399,10 @@
val [ordi] = goal Univ.thy "Ord(i) ==> ALL b. b : Vset(i) --> rank(b) < i";
by (rtac (ordi RS trans_induct) 1);
by (stac Vset 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (stac rank 1);
by (rtac UN_succ_least_lt 1);
-by (fast_tac ZF_cs 2);
+by (Fast_tac 2);
by (REPEAT (ares_tac [ltI] 1));
qed "Vset_rank_imp1";
@@ -413,7 +413,7 @@
by (rtac (ordi RS trans_induct) 1);
by (rtac allI 1);
by (stac Vset 1);
-by (fast_tac (ZF_cs addSIs [rank_lt RS ltD]) 1);
+by (fast_tac (!claset addSIs [rank_lt RS ltD]) 1);
qed "Vset_rank_imp2";
goal Univ.thy "!!x i. rank(x)<i ==> x : Vset(i)";
@@ -435,7 +435,7 @@
goal Univ.thy "!!i. Ord(i) ==> rank(Vset(i)) = i";
by (stac rank 1);
by (rtac equalityI 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (EVERY' [rtac UN_I,
etac (i_subset_Vfrom RS subsetD),
etac (Ord_in_Ord RS rank_of_Ord RS ssubst),
@@ -463,16 +463,14 @@
val rank_rls = [rank_Inl, rank_Inr, rank_pair1, rank_pair2];
val rank_trans_rls = rank_rls @ (rank_rls RLN (2, [lt_trans]));
-val rank_ss = ZF_ss
- addsimps [case_Inl, case_Inr, VsetI]
- addsimps rank_trans_rls;
+val rank_ss = !simpset addsimps [VsetI] addsimps rank_trans_rls;
(** Recursion over Vset levels! **)
(*NOT SUITABLE FOR REWRITING: recursive!*)
goalw Univ.thy [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))";
by (stac transrec 1);
-by (simp_tac (ZF_ss addsimps [Ord_rank, Ord_succ, VsetD RS ltD RS beta,
+by (simp_tac (!simpset addsimps [Ord_rank, Ord_succ, VsetD RS ltD RS beta,
VsetI RS beta, le_refl]) 1);
qed "Vrec";
@@ -581,7 +579,7 @@
qed "two_in_univ";
goalw Univ.thy [bool_def] "bool <= univ(A)";
-by (fast_tac (ZF_cs addSIs [zero_in_univ,one_in_univ]) 1);
+by (fast_tac (!claset addSIs [zero_in_univ,one_in_univ]) 1);
qed "bool_subset_univ";
bind_thm ("bool_into_univ", (bool_subset_univ RS subsetD));
@@ -616,21 +614,21 @@
goal Univ.thy
"!!i. [| b: Fin(Vfrom(A,i)); Limit(i) |] ==> EX j. b <= Vfrom(A,j) & j<i";
by (etac Fin_induct 1);
-by (fast_tac (ZF_cs addSDs [Limit_has_0]) 1);
-by (safe_tac ZF_cs);
+by (fast_tac (!claset addSDs [Limit_has_0]) 1);
+by (safe_tac (!claset));
by (etac Limit_VfromE 1);
by (assume_tac 1);
by (res_inst_tac [("x", "xa Un j")] exI 1);
-by (best_tac (ZF_cs addIs [subset_refl RS Vfrom_mono RS subsetD,
+by (best_tac (!claset addIs [subset_refl RS Vfrom_mono RS subsetD,
Un_least_lt]) 1);
val Fin_Vfrom_lemma = result();
goal Univ.thy "!!i. Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)";
by (rtac subsetI 1);
by (dtac Fin_Vfrom_lemma 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (resolve_tac [Vfrom RS ssubst] 1);
-by (fast_tac (ZF_cs addSDs [ltD]) 1);
+by (fast_tac (!claset addSDs [ltD]) 1);
val Fin_VLimit = result();
bind_thm ("Fin_subset_VLimit", [Fin_mono, Fin_VLimit] MRS subset_trans);
--- a/src/ZF/Univ.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Univ.thy Fri Jan 03 15:01:55 1997 +0100
@@ -11,7 +11,7 @@
But Ind_Syntax.univ refers to the constant "univ"
*)
-Univ = Arith + Sum + Finite + "mono" +
+Univ = Arith + Sum + Finite + mono +
consts
Vfrom :: [i,i]=>i
Vset :: i=>i
--- a/src/ZF/WF.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/WF.ML Fri Jan 03 15:01:55 1997 +0100
@@ -25,23 +25,23 @@
(** Equivalences between wf and wf_on **)
goalw WF.thy [wf_def, wf_on_def] "!!A r. wf(r) ==> wf[A](r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_imp_wf_on";
goalw WF.thy [wf_def, wf_on_def] "!!r. wf[field(r)](r) ==> wf(r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_on_field_imp_wf";
goal WF.thy "wf(r) <-> wf[field(r)](r)";
-by (fast_tac (ZF_cs addSEs [wf_imp_wf_on, wf_on_field_imp_wf]) 1);
+by (fast_tac (!claset addSEs [wf_imp_wf_on, wf_on_field_imp_wf]) 1);
qed "wf_iff_wf_on_field";
goalw WF.thy [wf_on_def, wf_def] "!!A B r. [| wf[A](r); B<=A |] ==> wf[B](r)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_on_subset_A";
goalw WF.thy [wf_on_def, wf_def] "!!A r s. [| wf[A](r); s<=r |] ==> wf[A](s)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_on_subset_r";
(** Introduction rules for wf_on **)
@@ -52,7 +52,7 @@
\ ==> wf[A](r)";
by (rtac (equals0I RS disjCI RS allI) 1);
by (res_inst_tac [ ("Z", "Z") ] prem 1);
-by (ALLGOALS (fast_tac ZF_cs));
+by (ALLGOALS (Fast_tac));
qed "wf_onI";
(*If r allows well-founded induction over A then wf[A](r)
@@ -65,8 +65,8 @@
by (rtac wf_onI 1);
by (res_inst_tac [ ("c", "u") ] (prem RS DiffE) 1);
by (contr_tac 3);
-by (fast_tac ZF_cs 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 2);
+by (Fast_tac 1);
qed "wf_onI2";
@@ -79,11 +79,11 @@
\ |] ==> P(a)";
by (res_inst_tac [ ("x", "{z:domain(r) Un {a}. ~P(z)}") ] (major RS allE) 1);
by (etac disjE 1);
-by (fast_tac (ZF_cs addEs [equalityE]) 1);
-by (asm_full_simp_tac (ZF_ss addsimps [domainI]) 1);
+by (fast_tac (!claset addEs [equalityE]) 1);
+by (asm_full_simp_tac (!simpset addsimps [domainI]) 1);
by (etac bexE 1);
by (dtac minor 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_induct";
(*Perform induction on i, then prove the wf(r) subgoal using prems. *)
@@ -101,11 +101,11 @@
by (wf_ind_tac "a" [wfr] 1);
by (rtac impI 1);
by (eresolve_tac prems 1);
-by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1);
+by (fast_tac (!claset addIs (prems RL [subsetD])) 1);
qed "wf_induct2";
-goal ZF.thy "!!r A. field(r Int A*A) <= A";
-by (fast_tac ZF_cs 1);
+goal domrange.thy "!!r A. field(r Int A*A) <= A";
+by (Fast_tac 1);
qed "field_Int_square";
val wfr::amem::prems = goalw WF.thy [wf_on_def]
@@ -114,7 +114,7 @@
\ |] ==> P(a)";
by (rtac ([wfr, amem, field_Int_square] MRS wf_induct2) 1);
by (REPEAT (ares_tac prems 1));
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_on_induct";
fun wf_on_ind_tac a prems i =
@@ -137,26 +137,26 @@
goal WF.thy "!!r. wf(r) ==> <a,a> ~: r";
by (wf_ind_tac "a" [] 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_not_refl";
goal WF.thy "!!r. [| wf(r); <a,x>:r; <x,a>:r |] ==> P";
by (subgoal_tac "ALL x. <a,x>:r --> <x,a>:r --> P" 1);
by (wf_ind_tac "a" [] 2);
-by (fast_tac ZF_cs 2);
-by (fast_tac FOL_cs 1);
+by (Fast_tac 2);
+by (Fast_tac 1);
qed "wf_asym";
goal WF.thy "!!r. [| wf[A](r); a: A |] ==> <a,a> ~: r";
by (wf_on_ind_tac "a" [] 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_on_not_refl";
goal WF.thy "!!r. [| wf[A](r); <a,b>:r; <b,a>:r; a:A; b:A |] ==> P";
by (subgoal_tac "ALL y:A. <a,y>:r --> <y,a>:r --> P" 1);
by (wf_on_ind_tac "a" [] 2);
-by (fast_tac ZF_cs 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 2);
+by (Fast_tac 1);
qed "wf_on_asym";
(*Needed to prove well_ordI. Could also reason that wf[A](r) means
@@ -166,8 +166,8 @@
by (subgoal_tac
"ALL y:A. ALL z:A. <a,y>:r --> <y,z>:r --> <z,a>:r --> P" 1);
by (wf_on_ind_tac "a" [] 2);
-by (fast_tac ZF_cs 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 2);
+by (Fast_tac 1);
qed "wf_on_chain3";
@@ -183,17 +183,17 @@
by (rtac (impI RS ballI) 1);
by (etac tranclE 1);
by (etac (bspec RS mp) 1 THEN assume_tac 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (cut_facts_tac [subs] 1);
(*astar_tac is slightly faster*)
-by (best_tac ZF_cs 1);
+by (Best_tac 1);
qed "wf_on_trancl";
goal WF.thy "!!r. wf(r) ==> wf(r^+)";
-by (asm_full_simp_tac (ZF_ss addsimps [wf_iff_wf_on_field]) 1);
+by (asm_full_simp_tac (!simpset addsimps [wf_iff_wf_on_field]) 1);
by (rtac (trancl_type RS field_rel_subset RSN (2, wf_on_subset_A)) 1);
by (etac wf_on_trancl 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_trancl";
@@ -227,7 +227,7 @@
eresolve_tac [underD, transD, spec RS mp]));
(*** NOTE! some simplifications need a different solver!! ***)
-val wf_super_ss = ZF_ss setsolver indhyp_tac;
+val wf_super_ss = !simpset setsolver indhyp_tac;
val prems = goalw WF.thy [is_recfun_def]
"[| wf(r); trans(r); is_recfun(r,a,H,f); is_recfun(r,b,H,g) |] ==> \
@@ -302,7 +302,7 @@
\ wftrec(r,a,H) = H(a, lam x: r-``{a}. wftrec(r,x,H))";
by (stac (rewrite_rule [is_recfun_def] unfold_the_recfun) 1);
by (ALLGOALS (asm_simp_tac
- (ZF_ss addsimps [vimage_singleton_iff RS iff_sym, the_recfun_cut])));
+ (!simpset addsimps [vimage_singleton_iff RS iff_sym, the_recfun_cut])));
qed "wftrec";
(** Removal of the premise trans(r) **)
@@ -340,6 +340,6 @@
"!!A r. [| wf[A](r); a: A |] ==> \
\ wfrec[A](r,a,H) = H(a, lam x: (r-``{a}) Int A. wfrec[A](r,x,H))";
by (etac (wfrec RS trans) 1);
-by (asm_simp_tac (ZF_ss addsimps [vimage_Int_square, cons_subset_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [vimage_Int_square, cons_subset_iff]) 1);
qed "wfrec_on";
--- a/src/ZF/WF.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/WF.thy Fri Jan 03 15:01:55 1997 +0100
@@ -6,7 +6,7 @@
Well-founded Recursion
*)
-WF = Trancl + "mono" + "equalities" +
+WF = Trancl + mono + equalities +
consts
wf :: i=>o
wf_on :: [i,i]=>o ("wf[_]'(_')")
--- a/src/ZF/ZF.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ZF.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,27 +8,13 @@
open ZF;
-exception CS_DATA of claset;
-
-let fun merge [] = CS_DATA empty_cs
- | merge cs = let val cs = map (fn CS_DATA x => x) cs;
- in CS_DATA (foldl merge_cs (hd cs, tl cs)) end;
-
- fun put (CS_DATA cs) = claset := cs;
-
- fun get () = CS_DATA (!claset);
-in add_thydata "ZF"
- ("claset", ThyMethods {merge = merge, put = put, get = get})
-end;
-
-claset:=empty_cs;
-
(*Useful examples: singletonI RS subst_elem, subst_elem RSN (2,IntI) *)
goal ZF.thy "!!a b A. [| b:A; a=b |] ==> a:A";
by (etac ssubst 1);
by (assume_tac 1);
val subst_elem = result();
+
(*** Bounded universal quantifier ***)
qed_goalw "ballI" ZF.thy [Ball_def]
@@ -60,12 +46,17 @@
[ (rtac (major RS ballE) 1),
(REPEAT (eresolve_tac prems 1)) ]);
+AddSIs [ballI];
+AddEs [rev_ballE];
+
(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*)
val ball_tac = dtac bspec THEN' assume_tac;
(*Trival rewrite rule; (ALL x:A.P)<->P holds only if A is nonempty!*)
qed_goal "ball_simp" ZF.thy "(ALL x:A. True) <-> True"
- (fn _=> [ (REPEAT (ares_tac [TrueI,ballI,iffI] 1)) ]);
+ (fn _=> [ Fast_tac 1 ]);
+
+Addsimps [ball_simp];
(*Congruence rule for rewriting*)
qed_goalw "ball_cong" ZF.thy [Ball_def]
@@ -75,8 +66,8 @@
(*** Bounded existential quantifier ***)
qed_goalw "bexI" ZF.thy [Bex_def]
- "[| P(x); x: A |] ==> EX x:A. P(x)"
- (fn prems=> [ (REPEAT (ares_tac (prems @ [exI,conjI]) 1)) ]);
+ "!!P. [| P(x); x: A |] ==> EX x:A. P(x)"
+ (fn _=> [ Fast_tac 1 ]);
(*Not of the general form for such rules; ~EX has become ALL~ *)
qed_goal "bexCI" ZF.thy
@@ -92,6 +83,9 @@
[ (rtac (major RS exE) 1),
(REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)) ]);
+AddIs [bexI];
+AddSEs [bexE];
+
(*We do not even have (EX x:A. True) <-> True unless A is nonempty!!*)
qed_goalw "bex_cong" ZF.thy [Bex_def]
@@ -99,6 +93,9 @@
\ |] ==> Bex(A,P) <-> Bex(A',P')"
(fn prems=> [ (simp_tac (FOL_ss addsimps prems addcongs [conj_cong]) 1) ]);
+Addcongs [ball_cong, bex_cong];
+
+
(*** Rules for subsets ***)
qed_goalw "subsetI" ZF.thy [subset_def]
@@ -118,24 +115,30 @@
[ (rtac (major RS ballE) 1),
(REPEAT (eresolve_tac prems 1)) ]);
+AddSIs [subsetI];
+AddEs [subsetCE, subsetD];
+
+
(*Takes assumptions A<=B; c:A and creates the assumption c:B *)
val set_mp_tac = dtac subsetD THEN' assume_tac;
(*Sometimes useful with premises in this order*)
qed_goal "rev_subsetD" ZF.thy "!!A B c. [| c:A; A<=B |] ==> c:B"
- (fn _=> [REPEAT (ares_tac [subsetD] 1)]);
+ (fn _=> [ Fast_tac 1 ]);
qed_goal "contra_subsetD" ZF.thy "!!c. [| A <= B; c ~: B |] ==> c ~: A"
- (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
+ (fn _=> [ Fast_tac 1 ]);
qed_goal "rev_contra_subsetD" ZF.thy "!!c. [| c ~: B; A <= B |] ==> c ~: A"
- (fn prems=> [ (REPEAT (eresolve_tac [asm_rl, contrapos, subsetD] 1)) ]);
+ (fn _=> [ Fast_tac 1 ]);
qed_goal "subset_refl" ZF.thy "A <= A"
- (fn _=> [ (rtac subsetI 1), atac 1 ]);
+ (fn _=> [ Fast_tac 1 ]);
-qed_goal "subset_trans" ZF.thy "[| A<=B; B<=C |] ==> A<=C"
- (fn prems=> [ (REPEAT (ares_tac ([subsetI]@(prems RL [subsetD])) 1)) ]);
+Addsimps [subset_refl];
+
+qed_goal "subset_trans" ZF.thy "!!A B C. [| A<=B; B<=C |] ==> A<=C"
+ (fn _=> [ Fast_tac 1 ]);
(*Useful for proving A<=B by rewriting in some cases*)
qed_goalw "subset_iff" ZF.thy [subset_def,Ball_def]
@@ -154,15 +157,8 @@
[ (rtac equalityI 1),
(REPEAT (ares_tac [subsetI, prem RS iffD1, prem RS iffD2] 1)) ]);
-qed_goal "equalityD1" ZF.thy "A = B ==> A<=B"
- (fn prems=>
- [ (rtac (extension RS iffD1 RS conjunct1) 1),
- (resolve_tac prems 1) ]);
-
-qed_goal "equalityD2" ZF.thy "A = B ==> B<=A"
- (fn prems=>
- [ (rtac (extension RS iffD1 RS conjunct2) 1),
- (resolve_tac prems 1) ]);
+bind_thm ("equalityD1", extension RS iffD1 RS conjunct1);
+bind_thm ("equalityD2", extension RS iffD1 RS conjunct2);
qed_goal "equalityE" ZF.thy
"[| A = B; [| A<=B; B<=A |] ==> P |] ==> P"
@@ -222,6 +218,9 @@
[ (rtac (major RS ReplaceE) 1),
(REPEAT (ares_tac prems 1)) ]);
+AddIs [ReplaceI];
+AddSEs [ReplaceE2];
+
qed_goal "Replace_cong" ZF.thy
"[| A=B; !!x y. x:B ==> P(x,y) <-> Q(x,y) |] ==> \
\ Replace(A,P) = Replace(B,Q)"
@@ -234,6 +233,8 @@
ORELSE (resolve_tac iffprems 1 THEN assume_tac 2))) ]
end);
+Addcongs [Replace_cong];
+
(*** Rules for RepFun ***)
qed_goalw "RepFunI" ZF.thy [RepFun_def]
@@ -253,28 +254,37 @@
[ (rtac (major RS ReplaceE) 1),
(REPEAT (ares_tac prems 1)) ]);
+AddIs [RepFunI];
+AddSEs [RepFunE];
+
qed_goalw "RepFun_cong" ZF.thy [RepFun_def]
"[| A=B; !!x. x:B ==> f(x)=g(x) |] ==> RepFun(A,f) = RepFun(B,g)"
- (fn prems=> [ (simp_tac (FOL_ss addcongs [Replace_cong] addsimps prems) 1) ]);
+ (fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);
+
+Addcongs [RepFun_cong];
qed_goalw "RepFun_iff" ZF.thy [Bex_def]
"b : {f(x). x:A} <-> (EX x:A. b=f(x))"
- (fn _ => [ (fast_tac (FOL_cs addIs [RepFunI] addSEs [RepFunE]) 1) ]);
+ (fn _ => [Fast_tac 1]);
+goal ZF.thy "{x.x:A} = A";
+by (fast_tac (!claset addSIs [equalityI]) 1);
+qed "triv_RepFun";
+
+Addsimps [RepFun_iff, triv_RepFun];
(*** Rules for Collect -- forming a subset by separation ***)
(*Separation is derivable from Replacement*)
qed_goalw "separation" ZF.thy [Collect_def]
"a : {x:A. P(x)} <-> a:A & P(a)"
- (fn _=> [ (fast_tac (FOL_cs addIs [bexI,ReplaceI]
- addSEs [bexE,ReplaceE]) 1) ]);
+ (fn _=> [Fast_tac 1]);
+
+Addsimps [separation];
qed_goal "CollectI" ZF.thy
- "[| a:A; P(a) |] ==> a : {x:A. P(x)}"
- (fn prems=>
- [ (rtac (separation RS iffD2) 1),
- (REPEAT (resolve_tac (prems@[conjI]) 1)) ]);
+ "!!P. [| a:A; P(a) |] ==> a : {x:A. P(x)}"
+ (fn _=> [ Asm_simp_tac 1 ]);
qed_goal "CollectE" ZF.thy
"[| a : {x:A. P(x)}; [| a:A; P(a) |] ==> R |] ==> R"
@@ -282,27 +292,27 @@
[ (rtac (separation RS iffD1 RS conjE) 1),
(REPEAT (ares_tac prems 1)) ]);
-qed_goal "CollectD1" ZF.thy "a : {x:A. P(x)} ==> a:A"
- (fn [major]=>
- [ (rtac (major RS CollectE) 1),
- (assume_tac 1) ]);
+qed_goal "CollectD1" ZF.thy "!!P. a : {x:A. P(x)} ==> a:A"
+ (fn _=> [ (etac CollectE 1), (assume_tac 1) ]);
-qed_goal "CollectD2" ZF.thy "a : {x:A. P(x)} ==> P(a)"
- (fn [major]=>
- [ (rtac (major RS CollectE) 1),
- (assume_tac 1) ]);
+qed_goal "CollectD2" ZF.thy "!!P. a : {x:A. P(x)} ==> P(a)"
+ (fn _=> [ (etac CollectE 1), (assume_tac 1) ]);
qed_goalw "Collect_cong" ZF.thy [Collect_def]
"[| A=B; !!x. x:B ==> P(x) <-> Q(x) |] ==> Collect(A,P) = Collect(B,Q)"
- (fn prems=> [ (simp_tac (FOL_ss addcongs [Replace_cong] addsimps prems) 1) ]);
+ (fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);
+
+AddSIs [CollectI];
+AddSEs [CollectE];
+Addcongs [Collect_cong];
(*** Rules for Unions ***)
+Addsimps [Union_iff];
+
(*The order of the premises presupposes that C is rigid; A may be flexible*)
-qed_goal "UnionI" ZF.thy "[| B: C; A: B |] ==> A: Union(C)"
- (fn prems=>
- [ (resolve_tac [Union_iff RS iffD2] 1),
- (REPEAT (resolve_tac (prems @ [bexI]) 1)) ]);
+qed_goal "UnionI" ZF.thy "!!C. [| B: C; A: B |] ==> A: Union(C)"
+ (fn _=> [ Simp_tac 1, Fast_tac 1 ]);
qed_goal "UnionE" ZF.thy
"[| A : Union(C); !!B.[| A: B; B: C |] ==> R |] ==> R"
@@ -310,47 +320,18 @@
[ (resolve_tac [Union_iff RS iffD1 RS bexE] 1),
(REPEAT (ares_tac prems 1)) ]);
-(*** Rules for Inter ***)
-
-(*Not obviously useful towards proving InterI, InterD, InterE*)
-qed_goalw "Inter_iff" ZF.thy [Inter_def,Ball_def]
- "A : Inter(C) <-> (ALL x:C. A: x) & (EX x. x:C)"
- (fn _=> [ (rtac (separation RS iff_trans) 1),
- (fast_tac (FOL_cs addIs [UnionI] addSEs [UnionE]) 1) ]);
-
-(* Intersection is well-behaved only if the family is non-empty! *)
-qed_goalw "InterI" ZF.thy [Inter_def]
- "[| !!x. x: C ==> A: x; c:C |] ==> A : Inter(C)"
- (fn prems=>
- [ (DEPTH_SOLVE (ares_tac ([CollectI,UnionI,ballI] @ prems) 1)) ]);
-
-(*A "destruct" rule -- every B in C contains A as an element, but
- A:B can hold when B:C does not! This rule is analogous to "spec". *)
-qed_goalw "InterD" ZF.thy [Inter_def]
- "[| A : Inter(C); B : C |] ==> A : B"
- (fn [major,minor]=>
- [ (rtac (major RS CollectD2 RS bspec) 1),
- (rtac minor 1) ]);
-
-(*"Classical" elimination rule -- does not require exhibiting B:C *)
-qed_goalw "InterE" ZF.thy [Inter_def]
- "[| A : Inter(C); A:B ==> R; B~:C ==> R |] ==> R"
- (fn major::prems=>
- [ (rtac (major RS CollectD2 RS ballE) 1),
- (REPEAT (eresolve_tac prems 1)) ]);
-
(*** Rules for Unions of families ***)
(* UN x:A. B(x) abbreviates Union({B(x). x:A}) *)
qed_goalw "UN_iff" ZF.thy [Bex_def]
"b : (UN x:A. B(x)) <-> (EX x:A. b : B(x))"
- (fn _=> [ (fast_tac (FOL_cs addIs [UnionI, RepFunI]
- addSEs [UnionE, RepFunE]) 1) ]);
+ (fn _=> [ Simp_tac 1, Fast_tac 1 ]);
+
+Addsimps [UN_iff];
(*The order of the premises presupposes that A is rigid; b may be flexible*)
-qed_goal "UN_I" ZF.thy "[| a: A; b: B(a) |] ==> b: (UN x:A. B(x))"
- (fn prems=>
- [ (REPEAT (resolve_tac (prems@[UnionI,RepFunI]) 1)) ]);
+qed_goal "UN_I" ZF.thy "!!A B. [| a: A; b: B(a) |] ==> b: (UN x:A. B(x))"
+ (fn _=> [ Simp_tac 1, Fast_tac 1 ]);
qed_goal "UN_E" ZF.thy
"[| b : (UN x:A. B(x)); !!x.[| x: A; b: B(x) |] ==> R |] ==> R"
@@ -360,23 +341,56 @@
qed_goal "UN_cong" ZF.thy
"[| A=B; !!x. x:B ==> C(x)=D(x) |] ==> (UN x:A.C(x)) = (UN x:B.D(x))"
- (fn prems=> [ (simp_tac (FOL_ss addcongs [RepFun_cong] addsimps prems) 1) ]);
+ (fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);
+
+(*No "Addcongs [UN_cong]" because UN is a combination of constants*)
+
+(* UN_E appears before UnionE so that it is tried first, to avoid expensive
+ calls to hyp_subst_tac. Cannot include UN_I as it is unsafe: would enlarge
+ the search space.*)
+AddIs [UnionI];
+AddSEs [UN_E];
+AddSEs [UnionE];
+
+
+(*** Rules for Inter ***)
+
+(*Not obviously useful towards proving InterI, InterD, InterE*)
+qed_goalw "Inter_iff" ZF.thy [Inter_def,Ball_def]
+ "A : Inter(C) <-> (ALL x:C. A: x) & (EX x. x:C)"
+ (fn _=> [ Simp_tac 1, Fast_tac 1 ]);
+(* Intersection is well-behaved only if the family is non-empty! *)
+qed_goalw "InterI" ZF.thy [Inter_def]
+ "[| !!x. x: C ==> A: x; c:C |] ==> A : Inter(C)"
+ (fn prems=> [ fast_tac (!claset addIs prems) 1 ]);
+
+(*A "destruct" rule -- every B in C contains A as an element, but
+ A:B can hold when B:C does not! This rule is analogous to "spec". *)
+qed_goalw "InterD" ZF.thy [Inter_def]
+ "!!C. [| A : Inter(C); B : C |] ==> A : B"
+ (fn _=> [ Fast_tac 1 ]);
+
+(*"Classical" elimination rule -- does not require exhibiting B:C *)
+qed_goalw "InterE" ZF.thy [Inter_def]
+ "[| A : Inter(C); A:B ==> R; B~:C ==> R |] ==> R"
+ (fn major::prems=>
+ [ (rtac (major RS CollectD2 RS ballE) 1),
+ (REPEAT (eresolve_tac prems 1)) ]);
+
+AddSIs [InterI];
+AddEs [InterD, make_elim InterD];
(*** Rules for Intersections of families ***)
(* INT x:A. B(x) abbreviates Inter({B(x). x:A}) *)
-qed_goal "INT_iff" ZF.thy
+qed_goalw "INT_iff" ZF.thy [Inter_def]
"b : (INT x:A. B(x)) <-> (ALL x:A. b : B(x)) & (EX x. x:A)"
- (fn _=> [ (simp_tac (FOL_ss addsimps [Inter_def, Ball_def, Bex_def,
- separation, Union_iff, RepFun_iff]) 1),
- (fast_tac FOL_cs 1) ]);
+ (fn _=> [ Simp_tac 1, Best_tac 1 ]);
qed_goal "INT_I" ZF.thy
"[| !!x. x: A ==> b: B(x); a: A |] ==> b: (INT x:A. B(x))"
- (fn prems=>
- [ (REPEAT (ares_tac (prems@[InterI,RepFunI]) 1
- ORELSE eresolve_tac [RepFunE,ssubst] 1)) ]);
+ (fn prems=> [ fast_tac (!claset addIs prems) 1 ]);
qed_goal "INT_E" ZF.thy
"[| b : (INT x:A. B(x)); a: A |] ==> b : B(a)"
@@ -386,7 +400,9 @@
qed_goal "INT_cong" ZF.thy
"[| A=B; !!x. x:B ==> C(x)=D(x) |] ==> (INT x:A.C(x)) = (INT x:B.D(x))"
- (fn prems=> [ (simp_tac (FOL_ss addcongs [RepFun_cong] addsimps prems) 1) ]);
+ (fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);
+
+(*No "Addcongs [INT_cong]" because INT is a combination of constants*)
(*** Rules for Powersets ***)
@@ -397,33 +413,38 @@
qed_goal "PowD" ZF.thy "A : Pow(B) ==> A<=B"
(fn [major]=> [ (rtac (major RS (Pow_iff RS iffD1)) 1) ]);
+AddSIs [PowI];
+AddSDs [PowD];
+
(*** Rules for the empty set ***)
(*The set {x:0.False} is empty; by foundation it equals 0
See Suppes, page 21.*)
-qed_goal "emptyE" ZF.thy "a:0 ==> P"
- (fn [major]=>
- [ (rtac (foundation RS disjE) 1),
- (etac (equalityD2 RS subsetD RS CollectD2 RS FalseE) 1),
- (rtac major 1),
- (etac bexE 1),
- (etac (CollectD2 RS FalseE) 1) ]);
+qed_goal "not_mem_empty" ZF.thy "a ~: 0"
+ (fn _=>
+ [ (cut_facts_tac [foundation] 1),
+ (best_tac (!claset addDs [equalityD2]) 1) ]);
+
+bind_thm ("emptyE", not_mem_empty RS notE);
+
+Addsimps [not_mem_empty];
+AddSEs [emptyE];
qed_goal "empty_subsetI" ZF.thy "0 <= A"
- (fn _ => [ (REPEAT (ares_tac [equalityI,subsetI,emptyE] 1)) ]);
+ (fn _=> [ Fast_tac 1 ]);
+
+Addsimps [empty_subsetI];
+AddSIs [empty_subsetI];
qed_goal "equals0I" ZF.thy "[| !!y. y:A ==> False |] ==> A=0"
- (fn prems=>
- [ (REPEAT (ares_tac (prems@[empty_subsetI,subsetI,equalityI]) 1
- ORELSE eresolve_tac (prems RL [FalseE]) 1)) ]);
+ (fn prems=> [ fast_tac (!claset addIs [equalityI] addDs prems) 1 ]);
-qed_goal "equals0D" ZF.thy "[| A=0; a:A |] ==> P"
- (fn [major,minor]=>
- [ (rtac (minor RS (major RS equalityD1 RS subsetD RS emptyE)) 1) ]);
+qed_goal "equals0D" ZF.thy "!!P. [| A=0; a:A |] ==> P"
+ (fn _=> [ Full_simp_tac 1, Fast_tac 1 ]);
qed_goal "not_emptyI" ZF.thy "!!A a. a:A ==> A ~= 0"
- (fn _ => [REPEAT (ares_tac [notI, equals0D] 1)]);
+ (fn _=> [ Fast_tac 1 ]);
qed_goal "not_emptyE" ZF.thy "[| A ~= 0; !!x. x:A ==> R |] ==> R"
(fn [major,minor]=>
@@ -431,19 +452,6 @@
swap_res_tac [minor] 1,
assume_tac 1 ]);
-(*A claset that leaves <= formulae unchanged!
- UN_E appears before UnionE so that it is tried first, to avoid expensive
- calls to hyp_subst_tac. Cannot include UN_I as it is unsafe:
- would enlarge the search space.*)
-val subset0_cs = FOL_cs
- addSIs [ballI, InterI, CollectI, PowI, empty_subsetI]
- addIs [bexI, UnionI, ReplaceI, RepFunI]
- addSEs [bexE, make_elim PowD, UN_E, UnionE, ReplaceE2, RepFunE,
- CollectE, emptyE]
- addEs [rev_ballE, InterD, make_elim InterD, subsetD];
-
-(*Standard claset*)
-val lemmas_cs = subset0_cs addSIs [subsetI] addEs [subsetCE];
(*** Cantor's Theorem: There is no surjection from a set to its powerset. ***)
@@ -459,7 +467,5 @@
(*Lemma for the inductive definition in Zorn.thy*)
qed_goal "Union_in_Pow" ZF.thy
"!!Y. Y : Pow(Pow(A)) ==> Union(Y) : Pow(A)"
- (fn _ => [fast_tac lemmas_cs 1]);
+ (fn _ => [Fast_tac 1]);
-
-add_thy_reader_file "thy_data.ML";
--- a/src/ZF/ZF.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ZF.thy Fri Jan 03 15:01:55 1997 +0100
@@ -177,6 +177,7 @@
and Powerset Axioms using the following definitions. *)
Collect_def "Collect(A,P) == {y . x:A, x=y & P(x)}"
+ Inter_def "Inter(A) == { x:Union(A) . ALL y:A. x:y}"
(*Unordered pairs (Upair) express binary union/intersection and cons;
set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...)*)
@@ -184,19 +185,6 @@
Upair_def "Upair(a,b) == {y. x:Pow(Pow(0)), (x=0 & y=a) | (x=Pow(0) & y=b)}"
cons_def "cons(a,A) == Upair(a,a) Un A"
- (* Difference, general intersection, binary union and small intersection *)
-
- Diff_def "A - B == { x:A . ~(x:B) }"
- Inter_def "Inter(A) == { x:Union(A) . ALL y:A. x:y}"
- Un_def "A Un B == Union(Upair(A,B))"
- Int_def "A Int B == Inter(Upair(A,B))"
-
- (* Definite descriptions -- via Replace over the set "1" *)
-
- the_def "The(P) == Union({y . x:{0}, P(y)})"
- if_def "if(P,a,b) == THE z. P & z=a | ~P & z=b"
-
- (* Ordered pairs and disjoint union of a family of sets *)
(* this "symmetric" definition works better than {{a}, {a,b}} *)
Pair_def "<a,b> == {{a,a}, {a,b}}"
@@ -233,32 +221,9 @@
ML
(* Pattern-matching and 'Dependent' type operators *)
-(*
-local open Syntax
-fun pttrn s = const"@pttrn" $ s;
-fun pttrns s t = const"@pttrns" $ s $ t;
-
-fun split2(Abs(x,T,t)) =
- let val (pats,u) = split1 t
- in (pttrns (Free(x,T)) pats, subst_bounds([free x],u)) end
- | split2(Const("split",_) $ r) =
- let val (pats,s) = split2(r)
- val (pats2,t) = split1(s)
- in (pttrns (pttrn pats) pats2, t) end
-and split1(Abs(x,T,t)) = (Free(x,T), subst_bounds([free x],t))
- | split1(Const("split",_)$t) = split2(t);
-
-fun split_tr'(t::args) =
- let val (pats,ft) = split2(t)
- in list_comb(const"_lambda" $ pttrn pats $ ft, args) end;
-
-in
-*)
val print_translation =
[(*("split", split_tr'),*)
("Pi", dependent_tr' ("@PROD", "op ->")),
("Sigma", dependent_tr' ("@SUM", "op *"))];
-(*
-end;
-*)
+
--- a/src/ZF/Zorn.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/Zorn.ML Fri Jan 03 15:01:55 1997 +0100
@@ -14,12 +14,12 @@
(*** Section 1. Mathematical Preamble ***)
goal ZF.thy "!!A B C. (ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Union_lemma0";
goal ZF.thy
"!!A B C. [| c:C; ALL x:C. A<=x | x<=B |] ==> A<=Inter(C) | Inter(C)<=B";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Inter_lemma0";
@@ -51,7 +51,7 @@
\ !!Y. [| Y <= TFin(S,next); ALL y:Y. P(y) |] ==> P(Union(Y)) \
\ |] ==> P(n)";
by (rtac (major RS TFin.induct) 1);
-by (ALLGOALS (fast_tac (ZF_cs addIs prems)));
+by (ALLGOALS (fast_tac (!claset addIs prems)));
qed "TFin_induct";
(*Perform induction on n, then prove the major premise using prems. *)
@@ -73,7 +73,7 @@
\ |] ==> n<=m | next`m<=n";
by (cut_facts_tac prems 1);
by (rtac (major RS TFin_induct) 1);
-by (etac Union_lemma0 2); (*or just fast_tac ZF_cs*)
+by (etac Union_lemma0 2); (*or just Fast_tac*)
by (fast_tac (subset_cs addIs [increasing_trans]) 1);
qed "TFin_linear_lemma1";
@@ -131,7 +131,7 @@
by (rtac (major RS TFin_induct) 1);
by (dtac TFin_subsetD 1);
by (REPEAT (assume_tac 1));
-by (fast_tac (ZF_cs addEs [ssubst]) 1);
+by (fast_tac (!claset addEs [ssubst]) 1);
by (fast_tac (subset_cs addIs [TFin_is_subset]) 1);
qed "equal_next_upper";
@@ -173,7 +173,7 @@
\ |] ==> ch ` super(S,X) : super(S,X)";
by (etac apply_type 1);
by (rewrite_goals_tac [super_def, maxchain_def]);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "choice_super";
goal Zorn.thy
@@ -184,7 +184,7 @@
by (dtac choice_super 1);
by (assume_tac 1);
by (assume_tac 1);
-by (asm_full_simp_tac (ZF_ss addsimps [super_def]) 1);
+by (asm_full_simp_tac (!simpset addsimps [super_def]) 1);
qed "choice_not_equals";
(*This justifies Definition 4.4*)
@@ -199,18 +199,18 @@
by (rewtac increasing_def);
by (rtac CollectI 1);
by (rtac lam_type 1);
-by (asm_simp_tac (ZF_ss setloop split_tac [expand_if]) 1);
-by (fast_tac (ZF_cs addSIs [super_subset_chain RS subsetD,
+by (asm_simp_tac (!simpset setloop split_tac [expand_if]) 1);
+by (fast_tac (!claset addSIs [super_subset_chain RS subsetD,
chain_subset_Pow RS subsetD,
choice_super]) 1);
(*Now, verify that it increases*)
-by (asm_simp_tac (ZF_ss addsimps [Pow_iff, subset_refl]
+by (asm_simp_tac (!simpset addsimps [Pow_iff, subset_refl]
setloop split_tac [expand_if]) 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (dtac choice_super 1);
by (REPEAT (assume_tac 1));
by (rewtac super_def);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Hausdorff_next_exists";
(*Lemma 4*)
@@ -223,11 +223,11 @@
\ |] ==> c: chain(S)";
by (etac TFin_induct 1);
by (asm_simp_tac
- (ZF_ss addsimps [chain_subset_Pow RS subsetD,
+ (!simpset addsimps [chain_subset_Pow RS subsetD,
choice_super RS (super_subset_chain RS subsetD)]
setloop split_tac [expand_if]) 1);
by (rewtac chain_def);
-by (rtac CollectI 1 THEN fast_tac ZF_cs 1);
+by (rtac CollectI 1 THEN Fast_tac 1);
(*Cannot use safe_tac: the disjunction must be left alone*)
by (REPEAT (rtac ballI 1 ORELSE etac UnionE 1));
by (res_inst_tac [("m1","B"), ("n1","Ba")] (TFin_subset_linear RS disjE) 1);
@@ -251,7 +251,7 @@
by (assume_tac 2);
by (rtac refl 2);
by (asm_full_simp_tac
- (ZF_ss addsimps [subset_refl RS TFin_UnionI RS
+ (!simpset addsimps [subset_refl RS TFin_UnionI RS
(TFin.dom_subset RS subsetD)]
setloop split_tac [expand_if]) 1);
by (eresolve_tac [choice_not_equals RS notE] 1);
@@ -265,25 +265,25 @@
(*Used in the proof of Zorn's Lemma*)
goalw Zorn.thy [chain_def]
"!!c. [| c: chain(A); z: A; ALL x:c. x<=z |] ==> cons(z,c) : chain(A)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "chain_extend";
goal Zorn.thy
"!!S. ALL c: chain(S). Union(c) : S ==> EX y:S. ALL z:S. y<=z --> y=z";
by (resolve_tac [Hausdorff RS exE] 1);
-by (asm_full_simp_tac (ZF_ss addsimps [maxchain_def]) 1);
+by (asm_full_simp_tac (!simpset addsimps [maxchain_def]) 1);
by (rename_tac "c" 1);
by (res_inst_tac [("x", "Union(c)")] bexI 1);
-by (fast_tac ZF_cs 2);
-by (safe_tac ZF_cs);
+by (Fast_tac 2);
+by (safe_tac (!claset));
by (rename_tac "z" 1);
by (rtac classical 1);
by (subgoal_tac "cons(z,c): super(S,c)" 1);
-by (fast_tac (ZF_cs addEs [equalityE]) 1);
+by (fast_tac (!claset addEs [equalityE]) 1);
by (rewtac super_def);
by (safe_tac eq_cs);
-by (fast_tac (ZF_cs addEs [chain_extend]) 1);
-by (best_tac (ZF_cs addEs [equalityE]) 1);
+by (fast_tac (!claset addEs [chain_extend]) 1);
+by (best_tac (!claset addEs [equalityE]) 1);
qed "Zorn";
@@ -295,12 +295,12 @@
\ |] ==> ALL m:Z. n<=m";
by (cut_facts_tac prems 1);
by (rtac (major RS TFin_induct) 1);
-by (fast_tac ZF_cs 2); (*second induction step is easy*)
+by (Fast_tac 2); (*second induction step is easy*)
by (rtac ballI 1);
by (rtac (bspec RS TFin_subsetD RS disjE) 1);
by (REPEAT_SOME (eresolve_tac [asm_rl,subsetD]));
by (subgoal_tac "x = Inter(Z)" 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
by (fast_tac eq_cs 1);
qed "TFin_well_lemma5";
@@ -308,7 +308,7 @@
goal Zorn.thy "!!Z. [| Z <= TFin(S,next); z:Z |] ==> Inter(Z) : Z";
by (rtac classical 1);
by (subgoal_tac "Z = {Union(TFin(S,next))}" 1);
-by (asm_simp_tac (ZF_ss addsimps [Inter_singleton]) 1);
+by (asm_simp_tac (!simpset addsimps [Inter_singleton]) 1);
by (etac equal_singleton 1);
by (rtac (Union_upper RS equalityI) 1);
by (rtac (subset_refl RS TFin_UnionI RS TFin_well_lemma5 RS bspec) 2);
@@ -324,12 +324,12 @@
(*Prove the linearity goal first*)
by (REPEAT (rtac ballI 2));
by (excluded_middle_tac "x=y" 2);
-by (fast_tac ZF_cs 3);
+by (Fast_tac 3);
(*The x~=y case remains*)
by (res_inst_tac [("n1","x"), ("m1","y")]
(TFin_subset_linear RS disjE) 2 THEN REPEAT (assume_tac 2));
-by (fast_tac ZF_cs 2);
-by (fast_tac ZF_cs 2);
+by (Fast_tac 2);
+by (Fast_tac 2);
(*Now prove the well_foundedness goal*)
by (rtac wf_onI 1);
by (forward_tac [well_ord_TFin_lemma] 1 THEN assume_tac 1);
@@ -361,12 +361,12 @@
(*Verify that it increases*)
by (rtac allI 2);
by (rtac impI 2);
-by (asm_simp_tac (ZF_ss addsimps [Pow_iff, subset_consI, subset_refl]
+by (asm_simp_tac (!simpset addsimps [Pow_iff, subset_consI, subset_refl]
setloop split_tac [expand_if]) 2);
(*Type checking is surprisingly hard!*)
-by (asm_simp_tac (ZF_ss addsimps [Pow_iff, cons_subset_iff, subset_refl]
+by (asm_simp_tac (!simpset addsimps [Pow_iff, cons_subset_iff, subset_refl]
setloop split_tac [expand_if]) 1);
-by (fast_tac (ZF_cs addSIs [choice_Diff RS DiffD1]) 1);
+by (fast_tac (!claset addSIs [choice_Diff RS DiffD1]) 1);
qed "Zermelo_next_exists";
@@ -380,18 +380,19 @@
by (res_inst_tac [("d", "%y. ch`(S-y)")] lam_injective 1);
by (rtac DiffI 1);
by (resolve_tac [Collect_subset RS TFin_UnionI] 1);
-by (fast_tac (ZF_cs addSIs [Collect_subset RS TFin_UnionI]
- addEs [equalityE]) 1);
+by (fast_tac (!claset addSIs [Collect_subset RS TFin_UnionI]
+ addEs [equalityE]) 1);
by (subgoal_tac "x ~: Union({y: TFin(S,next). x~: y})" 1);
-by (fast_tac (ZF_cs addEs [equalityE]) 2);
+by (fast_tac (!claset addEs [equalityE]) 2);
by (subgoal_tac "Union({y: TFin(S,next). x~: y}) ~= S" 1);
-by (fast_tac (ZF_cs addEs [equalityE]) 2);
+by (fast_tac (!claset addEs [equalityE]) 2);
(*For proving x : next`Union(...);
Abrial & Laffitte's justification appears to be faulty.*)
by (subgoal_tac "~ next ` Union({y: TFin(S,next). x~: y}) <= \
\ Union({y: TFin(S,next). x~: y})" 1);
by (asm_simp_tac
- (ZF_ss addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
+ (!simpset delsimps [Union_iff]
+ addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
Pow_iff, cons_subset_iff, subset_refl,
choice_Diff RS DiffD2]
setloop split_tac [expand_if]) 2);
@@ -399,7 +400,7 @@
by (fast_tac (subset_cs addSIs [Collect_subset RS TFin_UnionI, TFin_nextI]) 2);
(*End of the lemmas!*)
by (asm_full_simp_tac
- (ZF_ss addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
+ (!simpset addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
Pow_iff, cons_subset_iff, subset_refl]
setloop split_tac [expand_if]) 1);
by (REPEAT (eresolve_tac [asm_rl, consE, sym, notE] 1));
--- a/src/ZF/domrange.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/domrange.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,19 +8,19 @@
(*** converse ***)
-qed_goalw "converse_iff" ZF.thy [converse_def]
+qed_goalw "converse_iff" thy [converse_def]
"<a,b>: converse(r) <-> <b,a>:r"
- (fn _ => [ (fast_tac pair_cs 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goalw "converseI" ZF.thy [converse_def]
+qed_goalw "converseI" thy [converse_def]
"!!a b r. <a,b>:r ==> <b,a>:converse(r)"
- (fn _ => [ (fast_tac pair_cs 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goalw "converseD" ZF.thy [converse_def]
+qed_goalw "converseD" thy [converse_def]
"!!a b r. <a,b> : converse(r) ==> <b,a> : r"
- (fn _ => [ (fast_tac pair_cs 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goalw "converseE" ZF.thy [converse_def]
+qed_goalw "converseE" thy [converse_def]
"[| yx : converse(r); \
\ !!x y. [| yx=<y,x>; <x,y>:r |] ==> P \
\ |] ==> P"
@@ -30,54 +30,62 @@
(hyp_subst_tac 1),
(assume_tac 1) ]);
-val converse_cs = pair_cs addSIs [converseI]
- addSEs [converseD,converseE];
+Addsimps [converse_iff];
+AddSIs [converseI];
+AddSEs [converseD,converseE];
-qed_goal "converse_converse" ZF.thy
+qed_goal "converse_converse" thy
"!!A B r. r<=Sigma(A,B) ==> converse(converse(r)) = r"
- (fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addSIs [equalityI]) 1) ]);
-qed_goal "converse_type" ZF.thy "!!A B r. r<=A*B ==> converse(r)<=B*A"
- (fn _ => [ (fast_tac converse_cs 1) ]);
+qed_goal "converse_type" thy "!!A B r. r<=A*B ==> converse(r)<=B*A"
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goal "converse_prod" ZF.thy "converse(A*B) = B*A"
- (fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]);
+qed_goal "converse_prod" thy "converse(A*B) = B*A"
+ (fn _ => [ (fast_tac (!claset addSIs [equalityI]) 1) ]);
-qed_goal "converse_empty" ZF.thy "converse(0) = 0"
- (fn _ => [ (fast_tac (converse_cs addSIs [equalityI]) 1) ]);
+qed_goal "converse_empty" thy "converse(0) = 0"
+ (fn _ => [ (fast_tac (!claset addSIs [equalityI]) 1) ]);
+
+Addsimps [converse_prod, converse_empty];
(*** domain ***)
-qed_goalw "domain_iff" ZF.thy [domain_def]
+qed_goalw "domain_iff" thy [domain_def]
"a: domain(r) <-> (EX y. <a,y>: r)"
- (fn _=> [ (fast_tac pair_cs 1) ]);
+ (fn _=> [ (Fast_tac 1) ]);
-qed_goal "domainI" ZF.thy "!!a b r. <a,b>: r ==> a: domain(r)"
- (fn _ => [ (etac (exI RS (domain_iff RS iffD2)) 1) ]);
+qed_goal "domainI" thy "!!a b r. <a,b>: r ==> a: domain(r)"
+ (fn _=> [ (etac (exI RS (domain_iff RS iffD2)) 1) ]);
-qed_goal "domainE" ZF.thy
+qed_goal "domainE" thy
"[| a : domain(r); !!y. <a,y>: r ==> P |] ==> P"
(fn prems=>
[ (rtac (domain_iff RS iffD1 RS exE) 1),
(REPEAT (ares_tac prems 1)) ]);
-qed_goal "domain_subset" ZF.thy "domain(Sigma(A,B)) <= A"
- (fn _ => [ (rtac subsetI 1), (etac domainE 1), (etac SigmaD1 1) ]);
+AddIs [domainI];
+AddSEs [domainE];
+qed_goal "domain_subset" thy "domain(Sigma(A,B)) <= A"
+ (fn _=> [ (Fast_tac 1) ]);
(*** range ***)
-qed_goalw "rangeI" ZF.thy [range_def] "!!a b r.<a,b>: r ==> b : range(r)"
+qed_goalw "rangeI" thy [range_def] "!!a b r.<a,b>: r ==> b : range(r)"
(fn _ => [ (etac (converseI RS domainI) 1) ]);
-qed_goalw "rangeE" ZF.thy [range_def]
+qed_goalw "rangeE" thy [range_def]
"[| b : range(r); !!x. <x,b>: r ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS domainE) 1),
(resolve_tac prems 1),
(etac converseD 1) ]);
-qed_goalw "range_subset" ZF.thy [range_def] "range(A*B) <= B"
+AddIs [rangeI];
+AddSEs [rangeE];
+
+qed_goalw "range_subset" thy [range_def] "range(A*B) <= B"
(fn _ =>
[ (stac converse_prod 1),
(rtac domain_subset 1) ]);
@@ -85,22 +93,17 @@
(*** field ***)
-qed_goalw "fieldI1" ZF.thy [field_def] "<a,b>: r ==> a : field(r)"
- (fn [prem]=>
- [ (rtac (prem RS domainI RS UnI1) 1) ]);
+qed_goalw "fieldI1" thy [field_def] "!!r. <a,b>: r ==> a : field(r)"
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goalw "fieldI2" ZF.thy [field_def] "<a,b>: r ==> b : field(r)"
- (fn [prem]=>
- [ (rtac (prem RS rangeI RS UnI2) 1) ]);
+qed_goalw "fieldI2" thy [field_def] "!!r. <a,b>: r ==> b : field(r)"
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goalw "fieldCI" ZF.thy [field_def]
+qed_goalw "fieldCI" thy [field_def]
"(~ <c,a>:r ==> <a,b>: r) ==> a : field(r)"
- (fn [prem]=>
- [ (rtac (prem RS domainI RS UnCI) 1),
- (swap_res_tac [rangeI] 1),
- (etac notnotD 1) ]);
+ (fn [prem]=> [ (fast_tac (!claset addIs [prem]) 1) ]);
-qed_goalw "fieldE" ZF.thy [field_def]
+qed_goalw "fieldE" thy [field_def]
"[| a : field(r); \
\ !!x. <a,x>: r ==> P; \
\ !!x. <x,a>: r ==> P |] ==> P"
@@ -108,96 +111,107 @@
[ (rtac (major RS UnE) 1),
(REPEAT (eresolve_tac (prems@[domainE,rangeE]) 1)) ]);
-qed_goal "field_subset" ZF.thy "field(A*B) <= A Un B"
- (fn _ => [ (fast_tac (pair_cs addIs [fieldCI] addSEs [fieldE]) 1) ]);
+AddIs [fieldCI];
+AddSEs [fieldE];
-qed_goalw "domain_subset_field" ZF.thy [field_def]
+qed_goal "field_subset" thy "field(A*B) <= A Un B"
+ (fn _ => [ (Fast_tac 1) ]);
+
+qed_goalw "domain_subset_field" thy [field_def]
"domain(r) <= field(r)"
(fn _ => [ (rtac Un_upper1 1) ]);
-qed_goalw "range_subset_field" ZF.thy [field_def]
+qed_goalw "range_subset_field" thy [field_def]
"range(r) <= field(r)"
(fn _ => [ (rtac Un_upper2 1) ]);
-qed_goal "domain_times_range" ZF.thy
+qed_goal "domain_times_range" thy
"!!A B r. r <= Sigma(A,B) ==> r <= domain(r)*range(r)"
- (fn _ => [ (fast_tac (pair_cs addIs [domainI,rangeI]) 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goal "field_times_field" ZF.thy
+qed_goal "field_times_field" thy
"!!A B r. r <= Sigma(A,B) ==> r <= field(r)*field(r)"
- (fn _ => [ (fast_tac (pair_cs addIs [fieldI1,fieldI2]) 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
(*** Image of a set under a function/relation ***)
-qed_goalw "image_iff" ZF.thy [image_def]
- "b : r``A <-> (EX x:A. <x,b>:r)"
- (fn _ => [ fast_tac (pair_cs addIs [rangeI]) 1 ]);
+qed_goalw "image_iff" thy [image_def] "b : r``A <-> (EX x:A. <x,b>:r)"
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goal "image_singleton_iff" ZF.thy
- "b : r``{a} <-> <a,b>:r"
+qed_goal "image_singleton_iff" thy "b : r``{a} <-> <a,b>:r"
(fn _ => [ rtac (image_iff RS iff_trans) 1,
- fast_tac pair_cs 1 ]);
+ Fast_tac 1 ]);
-qed_goalw "imageI" ZF.thy [image_def]
+qed_goalw "imageI" thy [image_def]
"!!a b r. [| <a,b>: r; a:A |] ==> b : r``A"
- (fn _ => [ (REPEAT (ares_tac [CollectI,rangeI,bexI] 1)) ]);
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goalw "imageE" ZF.thy [image_def]
+qed_goalw "imageE" thy [image_def]
"[| b: r``A; !!x.[| <x,b>: r; x:A |] ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS CollectE) 1),
(REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ]);
-qed_goal "image_subset" ZF.thy
- "!!A B r. r <= A*B ==> r``C <= B"
- (fn _ => [ (fast_tac (pair_cs addSEs [imageE]) 1) ]);
+AddIs [imageI];
+AddSEs [imageE];
+
+qed_goal "image_subset" thy "!!A B r. r <= A*B ==> r``C <= B"
+ (fn _ => [ (Fast_tac 1) ]);
(*** Inverse image of a set under a function/relation ***)
-qed_goalw "vimage_iff" ZF.thy [vimage_def,image_def,converse_def]
+qed_goalw "vimage_iff" thy [vimage_def,image_def,converse_def]
"a : r-``B <-> (EX y:B. <a,y>:r)"
- (fn _ => [ fast_tac (pair_cs addIs [rangeI]) 1 ]);
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goal "vimage_singleton_iff" ZF.thy
+qed_goal "vimage_singleton_iff" thy
"a : r-``{b} <-> <a,b>:r"
(fn _ => [ rtac (vimage_iff RS iff_trans) 1,
- fast_tac pair_cs 1 ]);
+ Fast_tac 1 ]);
-qed_goalw "vimageI" ZF.thy [vimage_def]
+qed_goalw "vimageI" thy [vimage_def]
"!!A B r. [| <a,b>: r; b:B |] ==> a : r-``B"
- (fn _ => [ (REPEAT (ares_tac [converseI RS imageI] 1)) ]);
+ (fn _ => [ (Fast_tac 1) ]);
-qed_goalw "vimageE" ZF.thy [vimage_def]
+qed_goalw "vimageE" thy [vimage_def]
"[| a: r-``B; !!x.[| <a,x>: r; x:B |] ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS imageE) 1),
(REPEAT (etac converseD 1 ORELSE ares_tac prems 1)) ]);
-qed_goalw "vimage_subset" ZF.thy [vimage_def]
+qed_goalw "vimage_subset" thy [vimage_def]
"!!A B r. r <= A*B ==> r-``C <= A"
(fn _ => [ (etac (converse_type RS image_subset) 1) ]);
(** Theorem-proving for ZF set theory **)
-val ZF_cs = pair_cs
- addSIs [converseI]
- addIs [imageI, vimageI, domainI, rangeI, fieldCI]
- addSEs [imageE, vimageE, domainE, rangeE, fieldE, converseD, converseE];
+AddIs [vimageI];
+AddSEs [vimageE];
+
+val ZF_cs = !claset;
val eq_cs = ZF_cs addSIs [equalityI];
(** The Union of a set of relations is a relation -- Lemma for fun_Union **)
-goal ZF.thy "!!S. (ALL x:S. EX A B. x <= A*B) ==> \
+goal thy "!!S. (ALL x:S. EX A B. x <= A*B) ==> \
\ Union(S) <= domain(Union(S)) * range(Union(S))";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "rel_Union";
(** The Union of 2 relations is a relation (Lemma for fun_Un) **)
-qed_goal "rel_Un" ZF.thy
+qed_goal "rel_Un" thy
"!!r s. [| r <= A*B; s <= C*D |] ==> (r Un s) <= (A Un C) * (B Un D)"
- (fn _ => [ (fast_tac ZF_cs 1) ]);
+ (fn _ => [ (Fast_tac 1) ]);
+goal thy "!!r. [| <a,c> : r; c~=b |] ==> domain(r-{<a,b>}) = domain(r)";
+by (deepen_tac eq_cs 0 1);
+qed "domain_Diff_eq";
+
+goal thy "!!r. [| <c,b> : r; c~=a |] ==> range(r-{<a,b>}) = range(r)";
+by (deepen_tac eq_cs 0 1);
+qed "range_Diff_eq";
+
--- a/src/ZF/domrange.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/domrange.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,4 @@
(*Dummy theory to document dependencies *)
-domrange = "pair" + "subset"
+domrange = pair + "subset"
+
--- a/src/ZF/equalities.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/equalities.ML Fri Jan 03 15:01:55 1997 +0100
@@ -10,23 +10,23 @@
(** Finite Sets **)
(* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*)
-goal ZF.thy "{a} Un B = cons(a,B)";
+goal thy "{a} Un B = cons(a,B)";
by (fast_tac eq_cs 1);
qed "cons_eq";
-goal ZF.thy "cons(a, cons(b, C)) = cons(b, cons(a, C))";
+goal thy "cons(a, cons(b, C)) = cons(b, cons(a, C))";
by (fast_tac eq_cs 1);
qed "cons_commute";
-goal ZF.thy "!!B. a: B ==> cons(a,B) = B";
+goal thy "!!B. a: B ==> cons(a,B) = B";
by (fast_tac eq_cs 1);
qed "cons_absorb";
-goal ZF.thy "!!B. a: B ==> cons(a, B-{a}) = B";
+goal thy "!!B. a: B ==> cons(a, B-{a}) = B";
by (fast_tac eq_cs 1);
qed "cons_Diff";
-goal ZF.thy "!!C. [| a: C; ALL y:C. y=b |] ==> C = {b}";
+goal thy "!!C. [| a: C; ALL y:C. y=b |] ==> C = {b}";
by (fast_tac eq_cs 1);
qed "equal_singleton_lemma";
val equal_singleton = ballI RSN (2,equal_singleton_lemma);
@@ -34,522 +34,523 @@
(** Binary Intersection **)
-goal ZF.thy "0 Int A = 0";
-by (fast_tac eq_cs 1);
-qed "Int_0";
-
(*NOT an equality, but it seems to belong here...*)
-goal ZF.thy "cons(a,B) Int C <= cons(a, B Int C)";
+goal thy "cons(a,B) Int C <= cons(a, B Int C)";
by (fast_tac eq_cs 1);
qed "Int_cons";
-goal ZF.thy "A Int A = A";
+goal thy "A Int A = A";
by (fast_tac eq_cs 1);
qed "Int_absorb";
-goal ZF.thy "A Int B = B Int A";
+goal thy "A Int B = B Int A";
by (fast_tac eq_cs 1);
qed "Int_commute";
-goal ZF.thy "(A Int B) Int C = A Int (B Int C)";
+goal thy "(A Int B) Int C = A Int (B Int C)";
by (fast_tac eq_cs 1);
qed "Int_assoc";
-goal ZF.thy "(A Un B) Int C = (A Int C) Un (B Int C)";
+goal thy "(A Un B) Int C = (A Int C) Un (B Int C)";
by (fast_tac eq_cs 1);
qed "Int_Un_distrib";
-goal ZF.thy "A<=B <-> A Int B = A";
+goal thy "A<=B <-> A Int B = A";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
qed "subset_Int_iff";
-goal ZF.thy "A<=B <-> B Int A = A";
+goal thy "A<=B <-> B Int A = A";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
qed "subset_Int_iff2";
-goal ZF.thy "!!A B C. C<=A ==> (A-B) Int C = C-B";
+goal thy "!!A B C. C<=A ==> (A-B) Int C = C-B";
by (fast_tac eq_cs 1);
qed "Int_Diff_eq";
(** Binary Union **)
-goal ZF.thy "0 Un A = A";
-by (fast_tac eq_cs 1);
-qed "Un_0";
-
-goal ZF.thy "cons(a,B) Un C = cons(a, B Un C)";
+goal thy "cons(a,B) Un C = cons(a, B Un C)";
by (fast_tac eq_cs 1);
qed "Un_cons";
-goal ZF.thy "A Un A = A";
+goal thy "A Un A = A";
by (fast_tac eq_cs 1);
qed "Un_absorb";
-goal ZF.thy "A Un B = B Un A";
+goal thy "A Un B = B Un A";
by (fast_tac eq_cs 1);
qed "Un_commute";
-goal ZF.thy "(A Un B) Un C = A Un (B Un C)";
+goal thy "(A Un B) Un C = A Un (B Un C)";
by (fast_tac eq_cs 1);
qed "Un_assoc";
-goal ZF.thy "(A Int B) Un C = (A Un C) Int (B Un C)";
+goal thy "(A Int B) Un C = (A Un C) Int (B Un C)";
by (fast_tac eq_cs 1);
qed "Un_Int_distrib";
-goal ZF.thy "A<=B <-> A Un B = B";
+goal thy "A<=B <-> A Un B = B";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
qed "subset_Un_iff";
-goal ZF.thy "A<=B <-> B Un A = B";
+goal thy "A<=B <-> B Un A = B";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
qed "subset_Un_iff2";
(** Simple properties of Diff -- set difference **)
-goal ZF.thy "A-A = 0";
+goal thy "A-A = 0";
by (fast_tac eq_cs 1);
qed "Diff_cancel";
-goal ZF.thy "0-A = 0";
+goal thy "0-A = 0";
by (fast_tac eq_cs 1);
qed "empty_Diff";
-goal ZF.thy "A-0 = A";
+goal thy "A-0 = A";
by (fast_tac eq_cs 1);
qed "Diff_0";
-goal ZF.thy "A-B=0 <-> A<=B";
+goal thy "A-B=0 <-> A<=B";
by (fast_tac (eq_cs addEs [equalityE]) 1);
qed "Diff_eq_0_iff";
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
-goal ZF.thy "A - cons(a,B) = A - B - {a}";
+goal thy "A - cons(a,B) = A - B - {a}";
by (fast_tac eq_cs 1);
qed "Diff_cons";
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*)
-goal ZF.thy "A - cons(a,B) = A - {a} - B";
+goal thy "A - cons(a,B) = A - {a} - B";
by (fast_tac eq_cs 1);
qed "Diff_cons2";
-goal ZF.thy "A Int (B-A) = 0";
+goal thy "A Int (B-A) = 0";
by (fast_tac eq_cs 1);
qed "Diff_disjoint";
-goal ZF.thy "!!A B. A<=B ==> A Un (B-A) = B";
+goal thy "!!A B. A<=B ==> A Un (B-A) = B";
by (fast_tac eq_cs 1);
qed "Diff_partition";
-goal ZF.thy "A <= B Un (A - B)";
-by (fast_tac ZF_cs 1);
+goal thy "A <= B Un (A - B)";
+by (Fast_tac 1);
qed "subset_Un_Diff";
-goal ZF.thy "!!A B. [| A<=B; B<=C |] ==> B-(C-A) = A";
+goal thy "!!A B. [| A<=B; B<=C |] ==> B-(C-A) = A";
by (fast_tac eq_cs 1);
qed "double_complement";
-goal ZF.thy "(A Un B) - (B-A) = A";
+goal thy "(A Un B) - (B-A) = A";
by (fast_tac eq_cs 1);
qed "double_complement_Un";
-goal ZF.thy
+goal thy
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)";
by (fast_tac eq_cs 1);
qed "Un_Int_crazy";
-goal ZF.thy "A - (B Un C) = (A-B) Int (A-C)";
+goal thy "A - (B Un C) = (A-B) Int (A-C)";
by (fast_tac eq_cs 1);
qed "Diff_Un";
-goal ZF.thy "A - (B Int C) = (A-B) Un (A-C)";
+goal thy "A - (B Int C) = (A-B) Un (A-C)";
by (fast_tac eq_cs 1);
qed "Diff_Int";
(*Halmos, Naive Set Theory, page 16.*)
-goal ZF.thy "(A Int B) Un C = A Int (B Un C) <-> C<=A";
+goal thy "(A Int B) Un C = A Int (B Un C) <-> C<=A";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
qed "Un_Int_assoc_iff";
(** Big Union and Intersection **)
-goal ZF.thy "Union(0) = 0";
-by (fast_tac eq_cs 1);
-qed "Union_0";
-
-goal ZF.thy "Union(cons(a,B)) = a Un Union(B)";
+goal thy "Union(cons(a,B)) = a Un Union(B)";
by (fast_tac eq_cs 1);
qed "Union_cons";
-goal ZF.thy "Union(A Un B) = Union(A) Un Union(B)";
+goal thy "Union(A Un B) = Union(A) Un Union(B)";
by (fast_tac eq_cs 1);
qed "Union_Un_distrib";
-goal ZF.thy "Union(A Int B) <= Union(A) Int Union(B)";
-by (fast_tac ZF_cs 1);
+goal thy "Union(A Int B) <= Union(A) Int Union(B)";
+by (Fast_tac 1);
qed "Union_Int_subset";
-goal ZF.thy "Union(C) Int A = 0 <-> (ALL B:C. B Int A = 0)";
+goal thy "Union(C) Int A = 0 <-> (ALL B:C. B Int A = 0)";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
qed "Union_disjoint";
-goalw ZF.thy [Inter_def] "Inter(0) = 0";
+goalw thy [Inter_def] "Inter(0) = 0";
by (fast_tac eq_cs 1);
qed "Inter_0";
-goal ZF.thy "!!A B. [| z:A; z:B |] ==> Inter(A) Un Inter(B) <= Inter(A Int B)";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. [| z:A; z:B |] ==> Inter(A) Un Inter(B) <= Inter(A Int B)";
+by (Fast_tac 1);
qed "Inter_Un_subset";
(* A good challenge: Inter is ill-behaved on the empty set *)
-goal ZF.thy "!!A B. [| a:A; b:B |] ==> Inter(A Un B) = Inter(A) Int Inter(B)";
+goal thy "!!A B. [| a:A; b:B |] ==> Inter(A Un B) = Inter(A) Int Inter(B)";
by (fast_tac eq_cs 1);
qed "Inter_Un_distrib";
-goal ZF.thy "Union({b}) = b";
+goal thy "Union({b}) = b";
by (fast_tac eq_cs 1);
qed "Union_singleton";
-goal ZF.thy "Inter({b}) = b";
+goal thy "Inter({b}) = b";
by (fast_tac eq_cs 1);
qed "Inter_singleton";
(** Unions and Intersections of Families **)
-goal ZF.thy "Union(A) = (UN x:A. x)";
+goal thy "Union(A) = (UN x:A. x)";
by (fast_tac eq_cs 1);
qed "Union_eq_UN";
-goalw ZF.thy [Inter_def] "Inter(A) = (INT x:A. x)";
+goalw thy [Inter_def] "Inter(A) = (INT x:A. x)";
by (fast_tac eq_cs 1);
qed "Inter_eq_INT";
-goal ZF.thy "(UN i:0. A(i)) = 0";
+goal thy "(UN i:0. A(i)) = 0";
by (fast_tac eq_cs 1);
qed "UN_0";
(*Halmos, Naive Set Theory, page 35.*)
-goal ZF.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))";
+goal thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))";
by (fast_tac eq_cs 1);
qed "Int_UN_distrib";
-goal ZF.thy "!!A B. i:I ==> B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))";
+goal thy "!!A B. i:I ==> B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))";
by (fast_tac eq_cs 1);
qed "Un_INT_distrib";
-goal ZF.thy
+goal thy
"(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))";
by (fast_tac eq_cs 1);
qed "Int_UN_distrib2";
-goal ZF.thy
+goal thy
"!!I J. [| i:I; j:J |] ==> \
\ (INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))";
by (fast_tac eq_cs 1);
qed "Un_INT_distrib2";
-goal ZF.thy "!!A. a: A ==> (UN y:A. c) = c";
+goal thy "!!A. a: A ==> (UN y:A. c) = c";
by (fast_tac eq_cs 1);
qed "UN_constant";
-goal ZF.thy "!!A. a: A ==> (INT y:A. c) = c";
+goal thy "!!A. a: A ==> (INT y:A. c) = c";
by (fast_tac eq_cs 1);
qed "INT_constant";
(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5:
Union of a family of unions **)
-goal ZF.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))";
+goal thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))";
by (fast_tac eq_cs 1);
qed "UN_Un_distrib";
-goal ZF.thy
+goal thy
"!!A B. i:I ==> \
\ (INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))";
by (fast_tac eq_cs 1);
qed "INT_Int_distrib";
-goal ZF.thy "(UN z:I Int J. A(z)) <= (UN z:I. A(z)) Int (UN z:J. A(z))";
-by (fast_tac ZF_cs 1);
+goal thy "(UN z:I Int J. A(z)) <= (UN z:I. A(z)) Int (UN z:J. A(z))";
+by (Fast_tac 1);
qed "UN_Int_subset";
(** Devlin, page 12, exercise 5: Complements **)
-goal ZF.thy "!!A B. i:I ==> B - (UN i:I. A(i)) = (INT i:I. B - A(i))";
+goal thy "!!A B. i:I ==> B - (UN i:I. A(i)) = (INT i:I. B - A(i))";
by (fast_tac eq_cs 1);
qed "Diff_UN";
-goal ZF.thy "!!A B. i:I ==> B - (INT i:I. A(i)) = (UN i:I. B - A(i))";
+goal thy "!!A B. i:I ==> B - (INT i:I. A(i)) = (UN i:I. B - A(i))";
by (fast_tac eq_cs 1);
qed "Diff_INT";
(** Unions and Intersections with General Sum **)
(*Not suitable for rewriting: LOOPS!*)
-goal ZF.thy "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)";
+goal thy "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)";
by (fast_tac eq_cs 1);
qed "Sigma_cons1";
(*Not suitable for rewriting: LOOPS!*)
-goal ZF.thy "A * cons(b,B) = A*{b} Un A*B";
+goal thy "A * cons(b,B) = A*{b} Un A*B";
by (fast_tac eq_cs 1);
qed "Sigma_cons2";
-goal ZF.thy "Sigma(succ(A), B) = ({A}*B(A)) Un Sigma(A,B)";
+goal thy "Sigma(succ(A), B) = ({A}*B(A)) Un Sigma(A,B)";
by (fast_tac eq_cs 1);
qed "Sigma_succ1";
-goal ZF.thy "A * succ(B) = A*{B} Un A*B";
+goal thy "A * succ(B) = A*{B} Un A*B";
by (fast_tac eq_cs 1);
qed "Sigma_succ2";
-goal ZF.thy "(SUM x:(UN y:A. C(y)). B(x)) = (UN y:A. SUM x:C(y). B(x))";
+goal thy "(SUM x:(UN y:A. C(y)). B(x)) = (UN y:A. SUM x:C(y). B(x))";
by (fast_tac eq_cs 1);
qed "SUM_UN_distrib1";
-goal ZF.thy "(SUM i:I. UN j:J. C(i,j)) = (UN j:J. SUM i:I. C(i,j))";
+goal thy "(SUM i:I. UN j:J. C(i,j)) = (UN j:J. SUM i:I. C(i,j))";
by (fast_tac eq_cs 1);
qed "SUM_UN_distrib2";
-goal ZF.thy "(SUM i:I Un J. C(i)) = (SUM i:I. C(i)) Un (SUM j:J. C(j))";
+goal thy "(SUM i:I Un J. C(i)) = (SUM i:I. C(i)) Un (SUM j:J. C(j))";
by (fast_tac eq_cs 1);
qed "SUM_Un_distrib1";
-goal ZF.thy "(SUM i:I. A(i) Un B(i)) = (SUM i:I. A(i)) Un (SUM i:I. B(i))";
+goal thy "(SUM i:I. A(i) Un B(i)) = (SUM i:I. A(i)) Un (SUM i:I. B(i))";
by (fast_tac eq_cs 1);
qed "SUM_Un_distrib2";
(*First-order version of the above, for rewriting*)
-goal ZF.thy "I * (A Un B) = I*A Un I*B";
+goal thy "I * (A Un B) = I*A Un I*B";
by (rtac SUM_Un_distrib2 1);
qed "prod_Un_distrib2";
-goal ZF.thy "(SUM i:I Int J. C(i)) = (SUM i:I. C(i)) Int (SUM j:J. C(j))";
+goal thy "(SUM i:I Int J. C(i)) = (SUM i:I. C(i)) Int (SUM j:J. C(j))";
by (fast_tac eq_cs 1);
qed "SUM_Int_distrib1";
-goal ZF.thy
+goal thy
"(SUM i:I. A(i) Int B(i)) = (SUM i:I. A(i)) Int (SUM i:I. B(i))";
by (fast_tac eq_cs 1);
qed "SUM_Int_distrib2";
(*First-order version of the above, for rewriting*)
-goal ZF.thy "I * (A Int B) = I*A Int I*B";
+goal thy "I * (A Int B) = I*A Int I*B";
by (rtac SUM_Int_distrib2 1);
qed "prod_Int_distrib2";
(*Cf Aczel, Non-Well-Founded Sets, page 115*)
-goal ZF.thy "(SUM i:I. A(i)) = (UN i:I. {i} * A(i))";
+goal thy "(SUM i:I. A(i)) = (UN i:I. {i} * A(i))";
by (fast_tac eq_cs 1);
qed "SUM_eq_UN";
(** Domain **)
-qed_goal "domain_of_prod" ZF.thy "!!A B. b:B ==> domain(A*B) = A"
+qed_goal "domain_of_prod" thy "!!A B. b:B ==> domain(A*B) = A"
(fn _ => [ fast_tac eq_cs 1 ]);
-qed_goal "domain_0" ZF.thy "domain(0) = 0"
+qed_goal "domain_0" thy "domain(0) = 0"
(fn _ => [ fast_tac eq_cs 1 ]);
-qed_goal "domain_cons" ZF.thy
+qed_goal "domain_cons" thy
"domain(cons(<a,b>,r)) = cons(a, domain(r))"
(fn _ => [ fast_tac eq_cs 1 ]);
-goal ZF.thy "domain(A Un B) = domain(A) Un domain(B)";
+goal thy "domain(A Un B) = domain(A) Un domain(B)";
by (fast_tac eq_cs 1);
qed "domain_Un_eq";
-goal ZF.thy "domain(A Int B) <= domain(A) Int domain(B)";
+goal thy "domain(A Int B) <= domain(A) Int domain(B)";
by (fast_tac eq_cs 1);
qed "domain_Int_subset";
-goal ZF.thy "domain(A) - domain(B) <= domain(A - B)";
+goal thy "domain(A) - domain(B) <= domain(A - B)";
by (fast_tac eq_cs 1);
qed "domain_Diff_subset";
-goal ZF.thy "domain(converse(r)) = range(r)";
+goal thy "domain(converse(r)) = range(r)";
by (fast_tac eq_cs 1);
qed "domain_converse";
+Addsimps [domain_0, domain_cons, domain_Un_eq, domain_converse];
(** Range **)
-qed_goal "range_of_prod" ZF.thy
+qed_goal "range_of_prod" thy
"!!a A B. a:A ==> range(A*B) = B"
(fn _ => [ fast_tac eq_cs 1 ]);
-qed_goal "range_0" ZF.thy "range(0) = 0"
+qed_goal "range_0" thy "range(0) = 0"
(fn _ => [ fast_tac eq_cs 1 ]);
-qed_goal "range_cons" ZF.thy
+qed_goal "range_cons" thy
"range(cons(<a,b>,r)) = cons(b, range(r))"
(fn _ => [ fast_tac eq_cs 1 ]);
-goal ZF.thy "range(A Un B) = range(A) Un range(B)";
+goal thy "range(A Un B) = range(A) Un range(B)";
by (fast_tac eq_cs 1);
qed "range_Un_eq";
-goal ZF.thy "range(A Int B) <= range(A) Int range(B)";
-by (fast_tac ZF_cs 1);
+goal thy "range(A Int B) <= range(A) Int range(B)";
+by (Fast_tac 1);
qed "range_Int_subset";
-goal ZF.thy "range(A) - range(B) <= range(A - B)";
+goal thy "range(A) - range(B) <= range(A - B)";
by (fast_tac eq_cs 1);
qed "range_Diff_subset";
-goal ZF.thy "range(converse(r)) = domain(r)";
+goal thy "range(converse(r)) = domain(r)";
by (fast_tac eq_cs 1);
qed "range_converse";
+Addsimps [range_0, range_cons, range_Un_eq, range_converse];
+
+
(** Field **)
-qed_goal "field_of_prod" ZF.thy "field(A*A) = A"
+qed_goal "field_of_prod" thy "field(A*A) = A"
(fn _ => [ fast_tac eq_cs 1 ]);
-qed_goal "field_0" ZF.thy "field(0) = 0"
+qed_goal "field_0" thy "field(0) = 0"
(fn _ => [ fast_tac eq_cs 1 ]);
-qed_goal "field_cons" ZF.thy
+qed_goal "field_cons" thy
"field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))"
- (fn _ => [ rtac equalityI 1, ALLGOALS (fast_tac ZF_cs) ]);
+ (fn _ => [ rtac equalityI 1, ALLGOALS (Fast_tac) ]);
-goal ZF.thy "field(A Un B) = field(A) Un field(B)";
+goal thy "field(A Un B) = field(A) Un field(B)";
by (fast_tac eq_cs 1);
qed "field_Un_eq";
-goal ZF.thy "field(A Int B) <= field(A) Int field(B)";
+goal thy "field(A Int B) <= field(A) Int field(B)";
by (fast_tac eq_cs 1);
qed "field_Int_subset";
-goal ZF.thy "field(A) - field(B) <= field(A - B)";
+goal thy "field(A) - field(B) <= field(A - B)";
by (fast_tac eq_cs 1);
qed "field_Diff_subset";
+goal thy "field(converse(r)) = field(r)";
+by (fast_tac eq_cs 1);
+qed "field_converse";
+
+Addsimps [field_0, field_cons, field_Un_eq, field_converse];
+
(** Image **)
-goal ZF.thy "r``0 = 0";
+goal thy "r``0 = 0";
by (fast_tac eq_cs 1);
qed "image_0";
-goal ZF.thy "r``(A Un B) = (r``A) Un (r``B)";
+goal thy "r``(A Un B) = (r``A) Un (r``B)";
by (fast_tac eq_cs 1);
qed "image_Un";
-goal ZF.thy "r``(A Int B) <= (r``A) Int (r``B)";
-by (fast_tac ZF_cs 1);
+goal thy "r``(A Int B) <= (r``A) Int (r``B)";
+by (Fast_tac 1);
qed "image_Int_subset";
-goal ZF.thy "(r Int A*A)``B <= (r``B) Int A";
-by (fast_tac ZF_cs 1);
+goal thy "(r Int A*A)``B <= (r``B) Int A";
+by (Fast_tac 1);
qed "image_Int_square_subset";
-goal ZF.thy "!!r. B<=A ==> (r Int A*A)``B = (r``B) Int A";
+goal thy "!!r. B<=A ==> (r Int A*A)``B = (r``B) Int A";
by (fast_tac eq_cs 1);
qed "image_Int_square";
+Addsimps [image_0, image_Un];
+
(** Inverse Image **)
-goal ZF.thy "r-``0 = 0";
+goal thy "r-``0 = 0";
by (fast_tac eq_cs 1);
qed "vimage_0";
-goal ZF.thy "r-``(A Un B) = (r-``A) Un (r-``B)";
+goal thy "r-``(A Un B) = (r-``A) Un (r-``B)";
by (fast_tac eq_cs 1);
qed "vimage_Un";
-goal ZF.thy "r-``(A Int B) <= (r-``A) Int (r-``B)";
-by (fast_tac ZF_cs 1);
+goal thy "r-``(A Int B) <= (r-``A) Int (r-``B)";
+by (Fast_tac 1);
qed "vimage_Int_subset";
-goal ZF.thy "(r Int A*A)-``B <= (r-``B) Int A";
-by (fast_tac ZF_cs 1);
+goal thy "(r Int A*A)-``B <= (r-``B) Int A";
+by (Fast_tac 1);
qed "vimage_Int_square_subset";
-goal ZF.thy "!!r. B<=A ==> (r Int A*A)-``B = (r-``B) Int A";
+goal thy "!!r. B<=A ==> (r Int A*A)-``B = (r-``B) Int A";
by (fast_tac eq_cs 1);
qed "vimage_Int_square";
+Addsimps [vimage_0, vimage_Un];
+
(** Converse **)
-goal ZF.thy "converse(A Un B) = converse(A) Un converse(B)";
+goal thy "converse(A Un B) = converse(A) Un converse(B)";
by (fast_tac eq_cs 1);
qed "converse_Un";
-goal ZF.thy "converse(A Int B) = converse(A) Int converse(B)";
+goal thy "converse(A Int B) = converse(A) Int converse(B)";
by (fast_tac eq_cs 1);
qed "converse_Int";
-goal ZF.thy "converse(A) - converse(B) = converse(A - B)";
+goal thy "converse(A - B) = converse(A) - converse(B)";
by (fast_tac eq_cs 1);
qed "converse_Diff";
-goal ZF.thy "converse(UN x:A. B(x)) = (UN x:A. converse(B(x)))";
+goal thy "converse(UN x:A. B(x)) = (UN x:A. converse(B(x)))";
by (fast_tac eq_cs 1);
qed "converse_UN";
(*Unfolding Inter avoids using excluded middle on A=0*)
-goalw ZF.thy [Inter_def] "converse(INT x:A. B(x)) = (INT x:A. converse(B(x)))";
+goalw thy [Inter_def] "converse(INT x:A. B(x)) = (INT x:A. converse(B(x)))";
by (fast_tac eq_cs 1);
qed "converse_INT";
+Addsimps [converse_Un, converse_Int, converse_Diff, converse_UN, converse_INT];
+
(** Pow **)
-goal ZF.thy "Pow(A) Un Pow(B) <= Pow(A Un B)";
-by (fast_tac upair_cs 1);
+goal thy "Pow(A) Un Pow(B) <= Pow(A Un B)";
+by (Fast_tac 1);
qed "Un_Pow_subset";
-goal ZF.thy "(UN x:A. Pow(B(x))) <= Pow(UN x:A. B(x))";
-by (fast_tac upair_cs 1);
+goal thy "(UN x:A. Pow(B(x))) <= Pow(UN x:A. B(x))";
+by (Fast_tac 1);
qed "UN_Pow_subset";
-goal ZF.thy "A <= Pow(Union(A))";
-by (fast_tac upair_cs 1);
+goal thy "A <= Pow(Union(A))";
+by (Fast_tac 1);
qed "subset_Pow_Union";
-goal ZF.thy "Union(Pow(A)) = A";
+goal thy "Union(Pow(A)) = A";
by (fast_tac eq_cs 1);
qed "Union_Pow_eq";
-goal ZF.thy "Pow(A) Int Pow(B) = Pow(A Int B)";
+goal thy "Pow(A Int B) = Pow(A) Int Pow(B)";
by (fast_tac eq_cs 1);
qed "Int_Pow_eq";
-goal ZF.thy "!!x A. x:A ==> (INT x:A. Pow(B(x))) = Pow(INT x:A. B(x))";
+goal thy "!!x A. x:A ==> (INT x:A. Pow(B(x))) = Pow(INT x:A. B(x))";
by (fast_tac eq_cs 1);
qed "INT_Pow_subset";
+Addsimps [Union_Pow_eq, Int_Pow_eq];
+
(** RepFun **)
-goal ZF.thy "{f(x).x:A}=0 <-> A=0";
+goal thy "{f(x).x:A}=0 <-> A=0";
by (fast_tac (eq_cs addSEs [equalityE]) 1);
qed "RepFun_eq_0_iff";
-goal ZF.thy "{f(x).x:0} = 0";
-by (fast_tac eq_cs 1);
-qed "RepFun_0";
-
(** Collect **)
-goal ZF.thy "Collect(A Un B, P) = Collect(A,P) Un Collect(B,P)";
+goal thy "Collect(A Un B, P) = Collect(A,P) Un Collect(B,P)";
by (fast_tac eq_cs 1);
qed "Collect_Un";
-goal ZF.thy "Collect(A Int B, P) = Collect(A,P) Int Collect(B,P)";
+goal thy "Collect(A Int B, P) = Collect(A,P) Int Collect(B,P)";
by (fast_tac eq_cs 1);
qed "Collect_Int";
-goal ZF.thy "Collect(A - B, P) = Collect(A,P) - Collect(B,P)";
+goal thy "Collect(A - B, P) = Collect(A,P) - Collect(B,P)";
by (fast_tac eq_cs 1);
qed "Collect_Diff";
-goal ZF.thy
- "{x:cons(a,B). P(x)} = if(P(a), cons(a, {x:B. P(x)}), {x:B. P(x)})";
-by (simp_tac (FOL_ss setloop split_tac [expand_if]) 1);
+goal thy "{x:cons(a,B). P(x)} = if(P(a), cons(a, {x:B. P(x)}), {x:B. P(x)})";
+by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
by (fast_tac eq_cs 1);
qed "Collect_cons";
--- a/src/ZF/equalities.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/equalities.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,4 @@
(*Dummy theory to document dependencies *)
-equalities = "domrange"
+equalities = domrange
+
--- a/src/ZF/ex/Acc.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Acc.ML Fri Jan 03 15:01:55 1997 +0100
@@ -17,12 +17,12 @@
(*The intended introduction rule*)
val prems = goal Acc.thy
"[| !!b. <b,a>:r ==> b: acc(r); a: field(r) |] ==> a: acc(r)";
-by (fast_tac (ZF_cs addIs (prems@acc.intrs)) 1);
+by (fast_tac (!claset addIs (prems@acc.intrs)) 1);
qed "accI";
goal Acc.thy "!!a b r. [| b: acc(r); <a,b>: r |] ==> a: acc(r)";
by (etac acc.elim 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "acc_downward";
val [major,indhyp] = goal Acc.thy
@@ -31,7 +31,7 @@
\ |] ==> P(a)";
by (rtac (major RS acc.induct) 1);
by (rtac indhyp 1);
-by (fast_tac ZF_cs 2);
+by (Fast_tac 2);
by (resolve_tac acc.intrs 1);
by (assume_tac 2);
by (etac (Collect_subset RS Pow_mono RS subsetD) 1);
@@ -40,7 +40,7 @@
goal Acc.thy "wf[acc(r)](r)";
by (rtac wf_onI2 1);
by (etac acc_induct 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "wf_on_acc";
(* field(r) <= acc(r) ==> wf(r) *)
@@ -52,7 +52,7 @@
by (rtac subset_refl 1);
by (rtac accI 1);
by (assume_tac 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "acc_wfD";
goal Acc.thy "wf(r) <-> field(r) <= acc(r)";
--- a/src/ZF/ex/BT.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/BT.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,6 +8,8 @@
open BT;
+Addsimps bt.case_eqns;
+
(*Perform induction on l, then prove the major premise using prems. *)
fun bt_ind_tac a prems i =
EVERY [res_inst_tac [("x",a)] bt.induct i,
@@ -26,7 +28,7 @@
goalw BT.thy (bt.defs@bt.con_defs) "bt(univ(A)) <= univ(A)";
by (rtac lfp_lowerbound 1);
by (rtac (A_subset_univ RS univ_mono) 2);
-by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
+by (fast_tac (!claset addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
Pair_in_univ]) 1);
qed "bt_univ";
@@ -45,7 +47,7 @@
goal BT.thy "bt_rec(Lf,c,h) = c";
by (rtac (bt_rec_def RS def_Vrec RS trans) 1);
-by (simp_tac (ZF_ss addsimps bt.case_eqns) 1);
+by (Simp_tac 1);
qed "bt_rec_Lf";
goal BT.thy
@@ -54,6 +56,8 @@
by (simp_tac (rank_ss addsimps (bt.case_eqns @ [rank_Br1, rank_Br2])) 1);
qed "bt_rec_Br";
+Addsimps [bt_rec_Lf, bt_rec_Br];
+
(*Type checking -- proved by induction, as usual*)
val prems = goal BT.thy
"[| t: bt(A); \
@@ -62,8 +66,7 @@
\ h(x,y,z,r,s): C(Br(x,y,z)) \
\ |] ==> bt_rec(t,c,h) : C(t)";
by (bt_ind_tac "t" prems 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps
- (prems@[bt_rec_Lf,bt_rec_Br]))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "bt_rec_type";
(** Versions for use with definitions **)
@@ -84,6 +87,7 @@
(** n_nodes **)
val [n_nodes_Lf,n_nodes_Br] = bt_recs n_nodes_def;
+Addsimps [n_nodes_Lf, n_nodes_Br];
val prems = goalw BT.thy [n_nodes_def]
"xs: bt(A) ==> n_nodes(xs) : nat";
@@ -94,6 +98,7 @@
(** n_leaves **)
val [n_leaves_Lf,n_leaves_Br] = bt_recs n_leaves_def;
+Addsimps [n_leaves_Lf, n_leaves_Br];
val prems = goalw BT.thy [n_leaves_def]
"xs: bt(A) ==> n_leaves(xs) : nat";
@@ -103,6 +108,7 @@
(** bt_reflect **)
val [bt_reflect_Lf, bt_reflect_Br] = bt_recs bt_reflect_def;
+Addsimps [bt_reflect_Lf, bt_reflect_Br];
goalw BT.thy [bt_reflect_def] "!!xs. xs: bt(A) ==> bt_reflect(xs) : bt(A)";
by (REPEAT (ares_tac (bt.intrs @ [bt_rec_type]) 1));
@@ -115,13 +121,7 @@
val bt_typechecks =
bt.intrs @ [bt_rec_type, n_nodes_type, n_leaves_type, bt_reflect_type];
-val bt_ss = arith_ss
- addsimps bt.case_eqns
- addsimps bt_typechecks
- addsimps [bt_rec_Lf, bt_rec_Br,
- n_nodes_Lf, n_nodes_Br,
- n_leaves_Lf, n_leaves_Br,
- bt_reflect_Lf, bt_reflect_Br];
+Addsimps bt_typechecks;
(*** theorems about n_leaves ***)
@@ -129,13 +129,13 @@
val prems = goal BT.thy
"t: bt(A) ==> n_leaves(bt_reflect(t)) = n_leaves(t)";
by (bt_ind_tac "t" prems 1);
-by (ALLGOALS (asm_simp_tac (bt_ss addsimps [add_commute, n_leaves_type])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [add_commute, n_leaves_type])));
qed "n_leaves_reflect";
val prems = goal BT.thy
"t: bt(A) ==> n_leaves(t) = succ(n_nodes(t))";
by (bt_ind_tac "t" prems 1);
-by (ALLGOALS (asm_simp_tac (bt_ss addsimps [add_succ_right])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [add_succ_right])));
qed "n_leaves_nodes";
(*** theorems about bt_reflect ***)
@@ -143,7 +143,7 @@
val prems = goal BT.thy
"t: bt(A) ==> bt_reflect(bt_reflect(t))=t";
by (bt_ind_tac "t" prems 1);
-by (ALLGOALS (asm_simp_tac bt_ss));
+by (ALLGOALS Asm_simp_tac);
qed "bt_reflect_bt_reflect_ident";
--- a/src/ZF/ex/Bin.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Bin.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,6 +8,8 @@
open Bin;
+Addsimps bin.case_eqns;
+
(*Perform induction on l, then prove the major premise using prems. *)
fun bin_ind_tac a prems i =
EVERY [res_inst_tac [("x",a)] bin.induct i,
@@ -43,7 +45,7 @@
\ |] ==> bin_rec(w,a,b,h) : C(w)";
by (bin_ind_tac "w" prems 1);
by (ALLGOALS
- (asm_simp_tac (ZF_ss addsimps (prems@[bin_rec_Plus, bin_rec_Minus,
+ (asm_simp_tac (!simpset addsimps (prems@[bin_rec_Plus, bin_rec_Minus,
bin_rec_Bcons]))));
qed "bin_rec_type";
@@ -71,24 +73,24 @@
([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]);
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus";
-by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
+by (Asm_simp_tac 1);
qed "norm_Bcons_Plus_0";
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,1) = Bcons(Plus,1)";
-by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
+by (Asm_simp_tac 1);
qed "norm_Bcons_Plus_1";
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,0) = Bcons(Minus,0)";
-by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
+by (Asm_simp_tac 1);
qed "norm_Bcons_Minus_0";
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,1) = Minus";
-by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
+by (Asm_simp_tac 1);
qed "norm_Bcons_Minus_1";
goalw Bin.thy [norm_Bcons_def]
"norm_Bcons(Bcons(w,x),b) = Bcons(Bcons(w,x),b)";
-by (asm_simp_tac (ZF_ss addsimps bin.case_eqns) 1);
+by (asm_simp_tac (!simpset addsimps bin.case_eqns) 1);
qed "norm_Bcons_Bcons";
val norm_Bcons_simps = [norm_Bcons_Plus_0, norm_Bcons_Plus_1,
@@ -108,7 +110,7 @@
goalw Bin.thy [norm_Bcons_def]
"!!w. [| w: bin; b: bool |] ==> norm_Bcons(w,b) : bin";
by (etac bin.elim 1);
-by (ALLGOALS (asm_simp_tac (ZF_ss addsimps bin.case_eqns)));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps bin.case_eqns)));
by (typechk_tac (bin_typechecks0@bool_typechecks));
qed "norm_Bcons_type";
@@ -143,10 +145,9 @@
[integ_of_bin_type, norm_Bcons_type, bin_succ_type, bin_pred_type,
bin_minus_type, bin_add_type, bin_mult_type];
-val bin_ss = integ_ss
- addsimps([bool_1I, bool_0I,
- bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @
- bin_recs integ_of_bin_def @ bool_simps @ bin_typechecks);
+Addsimps ([bool_1I, bool_0I,
+ bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @
+ bin_recs integ_of_bin_def @ bin_typechecks);
val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @
[bool_subset_nat RS subsetD];
@@ -159,7 +160,7 @@
"!!z v. [| z $+ v = z' $+ v'; \
\ z: integ; z': integ; v: integ; v': integ; w: integ |] \
\ ==> z $+ (v $+ w) = z' $+ (v' $+ w)";
-by (asm_simp_tac (integ_ss addsimps ([zadd_assoc RS sym])) 1);
+by (asm_simp_tac (!simpset addsimps ([zadd_assoc RS sym])) 1);
qed "zadd_assoc_cong";
goal Integ.thy
@@ -172,8 +173,7 @@
bind_thm ("zadd_assoc_znat", (znat_type RS zadd_assoc_swap));
-val carry_ss = bin_ss addsimps
- (bin_recs bin_succ_def @ bin_recs bin_pred_def);
+Addsimps (bin_recs bin_succ_def @ bin_recs bin_pred_def);
(*norm_Bcons preserves the integer value of its argument*)
@@ -181,29 +181,29 @@
"!!w. [| w: bin; b: bool |] ==> \
\ integ_of_bin(norm_Bcons(w,b)) = integ_of_bin(Bcons(w,b))";
by (etac bin.elim 1);
-by (asm_simp_tac (ZF_ss addsimps norm_Bcons_simps) 3);
+by (asm_simp_tac (!simpset addsimps norm_Bcons_simps) 3);
by (ALLGOALS (etac boolE));
-by (ALLGOALS (asm_simp_tac (bin_ss addsimps (norm_Bcons_simps))));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps (norm_Bcons_simps))));
qed "integ_of_bin_norm_Bcons";
goal Bin.thy
"!!w. w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)";
by (etac bin.induct 1);
-by (simp_tac carry_ss 1);
-by (simp_tac carry_ss 1);
+by (Simp_tac 1);
+by (Simp_tac 1);
by (etac boolE 1);
by (ALLGOALS
- (asm_simp_tac (carry_ss addsimps integ_of_bin_norm_Bcons::zadd_ac)));
+ (asm_simp_tac (!simpset addsimps integ_of_bin_norm_Bcons::zadd_ac)));
qed "integ_of_bin_succ";
goal Bin.thy
"!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)";
by (etac bin.induct 1);
-by (simp_tac carry_ss 1);
-by (simp_tac carry_ss 1);
+by (Simp_tac 1);
+by (Simp_tac 1);
by (etac boolE 1);
by (ALLGOALS
- (asm_simp_tac (carry_ss addsimps integ_of_bin_norm_Bcons::zadd_ac)));
+ (asm_simp_tac (!simpset addsimps integ_of_bin_norm_Bcons::zadd_ac)));
qed "integ_of_bin_pred";
(*These two results replace the definitions of bin_succ and bin_pred*)
@@ -211,69 +211,62 @@
(*** bin_minus: (unary!) negation of binary integers ***)
-val bin_minus_ss =
- bin_ss addsimps (bin_recs bin_minus_def @
- [integ_of_bin_succ, integ_of_bin_pred]);
+Addsimps (bin_recs bin_minus_def @
+ [integ_of_bin_succ, integ_of_bin_pred]);
goal Bin.thy
"!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)";
by (etac bin.induct 1);
-by (simp_tac bin_minus_ss 1);
-by (simp_tac bin_minus_ss 1);
+by (Simp_tac 1);
+by (Simp_tac 1);
by (etac boolE 1);
by (ALLGOALS
- (asm_simp_tac (bin_minus_ss addsimps [zminus_zadd_distrib, zadd_assoc])));
+ (asm_simp_tac (!simpset addsimps (zadd_ac@[zminus_zadd_distrib]))));
qed "integ_of_bin_minus";
(*** bin_add: binary addition ***)
goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
-by (asm_simp_tac bin_ss 1);
+by (Asm_simp_tac 1);
qed "bin_add_Plus";
goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
-by (asm_simp_tac bin_ss 1);
+by (Asm_simp_tac 1);
qed "bin_add_Minus";
goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Plus) = Bcons(v,x)";
-by (simp_tac bin_ss 1);
+by (Simp_tac 1);
qed "bin_add_Bcons_Plus";
goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Minus) = bin_pred(Bcons(v,x))";
-by (simp_tac bin_ss 1);
+by (Simp_tac 1);
qed "bin_add_Bcons_Minus";
goalw Bin.thy [bin_add_def]
"!!w y. [| w: bin; y: bool |] ==> \
\ bin_add(Bcons(v,x), Bcons(w,y)) = \
\ norm_Bcons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)";
-by (asm_simp_tac bin_ss 1);
+by (Asm_simp_tac 1);
qed "bin_add_Bcons_Bcons";
-val bin_add_simps = [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus,
- bin_add_Bcons_Minus, bin_add_Bcons_Bcons,
- integ_of_bin_succ, integ_of_bin_pred,
- integ_of_bin_norm_Bcons];
+Addsimps [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus,
+ bin_add_Bcons_Minus, bin_add_Bcons_Bcons,
+ integ_of_bin_succ, integ_of_bin_pred,
+ integ_of_bin_norm_Bcons];
-val bin_add_ss =
- bin_ss addsimps ([bool_subset_nat RS subsetD] @ bin_add_simps);
+Addsimps [bool_subset_nat RS subsetD];
goal Bin.thy
"!!v. v: bin ==> \
\ ALL w: bin. integ_of_bin(bin_add(v,w)) = \
\ integ_of_bin(v) $+ integ_of_bin(w)";
by (etac bin.induct 1);
-by (simp_tac bin_add_ss 1);
-by (simp_tac bin_add_ss 1);
+by (Simp_tac 1);
+by (Simp_tac 1);
by (rtac ballI 1);
by (bin_ind_tac "wa" [] 1);
-by (asm_simp_tac bin_add_ss 1);
-by (asm_simp_tac (bin_add_ss addsimps zadd_ac) 1);
-by (etac boolE 1);
-by (asm_simp_tac (bin_add_ss addsimps zadd_ac) 2);
-by (etac boolE 1);
-by (ALLGOALS (asm_simp_tac (bin_add_ss addsimps zadd_ac)));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps zadd_ac setloop (etac boolE))));
val integ_of_bin_add_lemma = result();
bind_thm("integ_of_bin_add", integ_of_bin_add_lemma RS bspec);
@@ -281,10 +274,9 @@
(*** bin_add: binary multiplication ***)
-val bin_mult_ss =
- bin_ss addsimps (bin_recs bin_mult_def @
- [integ_of_bin_minus, integ_of_bin_add,
- integ_of_bin_norm_Bcons]);
+Addsimps (bin_recs bin_mult_def @
+ [integ_of_bin_minus, integ_of_bin_add,
+ integ_of_bin_norm_Bcons]);
val major::prems = goal Bin.thy
"[| v: bin; w: bin |] ==> \
@@ -292,12 +284,12 @@
\ integ_of_bin(v) $* integ_of_bin(w)";
by (cut_facts_tac prems 1);
by (bin_ind_tac "v" [major] 1);
-by (asm_simp_tac bin_mult_ss 1);
-by (asm_simp_tac bin_mult_ss 1);
+by (Asm_simp_tac 1);
+by (Asm_simp_tac 1);
by (etac boolE 1);
-by (asm_simp_tac (bin_mult_ss addsimps [zadd_zmult_distrib]) 2);
+by (asm_simp_tac (!simpset addsimps [zadd_zmult_distrib]) 2);
by (asm_simp_tac
- (bin_mult_ss addsimps ([zadd_zmult_distrib, zmult_1] @ zadd_ac)) 1);
+ (!simpset addsimps ([zadd_zmult_distrib, zmult_1] @ zadd_ac)) 1);
qed "integ_of_bin_mult";
(**** Computations ****)
@@ -308,19 +300,19 @@
val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def;
goal Bin.thy "bin_succ(Bcons(w,1)) = Bcons(bin_succ(w), 0)";
-by (simp_tac carry_ss 1);
+by (Simp_tac 1);
qed "bin_succ_Bcons1";
goal Bin.thy "bin_succ(Bcons(w,0)) = norm_Bcons(w,1)";
-by (simp_tac carry_ss 1);
+by (Simp_tac 1);
qed "bin_succ_Bcons0";
goal Bin.thy "bin_pred(Bcons(w,1)) = norm_Bcons(w,0)";
-by (simp_tac carry_ss 1);
+by (Simp_tac 1);
qed "bin_pred_Bcons1";
goal Bin.thy "bin_pred(Bcons(w,0)) = Bcons(bin_pred(w), 1)";
-by (simp_tac carry_ss 1);
+by (Simp_tac 1);
qed "bin_pred_Bcons0";
(** extra rules for bin_minus **)
@@ -328,11 +320,11 @@
val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def;
goal Bin.thy "bin_minus(Bcons(w,1)) = bin_pred(Bcons(bin_minus(w), 0))";
-by (simp_tac bin_minus_ss 1);
+by (Simp_tac 1);
qed "bin_minus_Bcons1";
goal Bin.thy "bin_minus(Bcons(w,0)) = Bcons(bin_minus(w), 0)";
-by (simp_tac bin_minus_ss 1);
+by (Simp_tac 1);
qed "bin_minus_Bcons0";
(** extra rules for bin_add **)
@@ -340,19 +332,19 @@
goal Bin.thy
"!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,1)) = \
\ norm_Bcons(bin_add(v, bin_succ(w)), 0)";
-by (asm_simp_tac bin_add_ss 1);
+by (Asm_simp_tac 1);
qed "bin_add_Bcons_Bcons11";
goal Bin.thy
"!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,0)) = \
\ norm_Bcons(bin_add(v,w), 1)";
-by (asm_simp_tac bin_add_ss 1);
+by (Asm_simp_tac 1);
qed "bin_add_Bcons_Bcons10";
goal Bin.thy
"!!w y. [| w: bin; y: bool |] ==> \
\ bin_add(Bcons(v,0), Bcons(w,y)) = norm_Bcons(bin_add(v,w), y)";
-by (asm_simp_tac bin_add_ss 1);
+by (Asm_simp_tac 1);
qed "bin_add_Bcons_Bcons0";
(** extra rules for bin_mult **)
@@ -361,17 +353,17 @@
goal Bin.thy
"bin_mult(Bcons(v,1), w) = bin_add(norm_Bcons(bin_mult(v,w),0), w)";
-by (simp_tac bin_mult_ss 1);
+by (Simp_tac 1);
qed "bin_mult_Bcons1";
goal Bin.thy "bin_mult(Bcons(v,0), w) = norm_Bcons(bin_mult(v,w),0)";
-by (simp_tac bin_mult_ss 1);
+by (Simp_tac 1);
qed "bin_mult_Bcons0";
(*** The computation simpset ***)
-val bin_comp_ss = integ_ss
+val bin_comp_ss = simpset_of "Integ"
addsimps [integ_of_bin_add RS sym, (*invoke bin_add*)
integ_of_bin_minus RS sym, (*invoke bin_minus*)
integ_of_bin_mult RS sym, (*invoke bin_mult*)
@@ -389,6 +381,7 @@
norm_Bcons_simps
setsolver (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0));
+
(*** Examples of performing binary arithmetic by simplification ***)
proof_timing := true;
--- a/src/ZF/ex/Brouwer.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Brouwer.ML Fri Jan 03 15:01:55 1997 +0100
@@ -29,8 +29,8 @@
\ |] ==> P(b)";
by (rtac (major RS brouwer.induct) 1);
by (REPEAT_SOME (ares_tac prems));
-by (fast_tac (ZF_cs addEs [fun_weaken_type]) 1);
-by (fast_tac (ZF_cs addDs [apply_type]) 1);
+by (fast_tac (!claset addEs [fun_weaken_type]) 1);
+by (fast_tac (!claset addDs [apply_type]) 1);
qed "brouwer_induct2";
@@ -51,8 +51,8 @@
\ |] ==> P(w)";
by (rtac (major RS Well.induct) 1);
by (REPEAT_SOME (ares_tac prems));
-by (fast_tac (ZF_cs addEs [fun_weaken_type]) 1);
-by (fast_tac (ZF_cs addDs [apply_type]) 1);
+by (fast_tac (!claset addEs [fun_weaken_type]) 1);
+by (fast_tac (!claset addDs [apply_type]) 1);
qed "Well_induct2";
@@ -60,5 +60,5 @@
Well to prove this.*)
goal Brouwer.thy "Well(bool, %x.x) = 1 + (1 -> Well(bool, %x.x))";
by (resolve_tac [Well_unfold RS trans] 1);
-by (simp_tac (ZF_ss addsimps [Sigma_bool, Pi_empty1, succ_def]) 1);
+by (simp_tac (!simpset addsimps [Sigma_bool, Pi_empty1, succ_def]) 1);
qed "Well_bool_unfold";
--- a/src/ZF/ex/CoUnit.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/CoUnit.ML Fri Jan 03 15:01:55 1997 +0100
@@ -25,7 +25,7 @@
by (etac counit.coinduct 1);
by (rtac subset_refl 1);
by (rewrite_goals_tac counit.con_defs);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "counit_eq_univ";
@@ -47,7 +47,7 @@
by (rtac (singletonI RS counit2.coinduct) 1);
by (rtac (qunivI RS singleton_subsetI) 1);
by (rtac ([lfp_subset, empty_subsetI RS univ_mono] MRS subset_trans) 1);
-by (fast_tac (ZF_cs addSIs [Con2_bnd_mono RS lfp_Tarski]) 1);
+by (fast_tac (!claset addSIs [Con2_bnd_mono RS lfp_Tarski]) 1);
qed "lfp_Con2_in_counit2";
(*Lemma for proving finality. Borrowed from ex/llist_eq.ML!*)
@@ -58,10 +58,9 @@
by (etac counit2.elim 1);
by (etac counit2.elim 1);
by (rewrite_goals_tac counit2.con_defs);
-val lleq_cs = subset_cs
- addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono]
- addSEs [Ord_in_Ord, Pair_inject];
-by (fast_tac lleq_cs 1);
+by (fast_tac (subset_cs
+ addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono]
+ addSEs [Ord_in_Ord, Pair_inject]) 1);
qed "counit2_Int_Vset_subset_lemma";
val counit2_Int_Vset_subset = standard
--- a/src/ZF/ex/Comb.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Comb.ML Fri Jan 03 15:01:55 1997 +0100
@@ -29,7 +29,7 @@
val contract_combD2 = contract.dom_subset RS subsetD RS SigmaD2;
goal Comb.thy "field(contract) = comb";
-by (fast_tac (ZF_cs addIs [equalityI,contract.K] addSEs [contract_combE2]) 1);
+by (fast_tac (!claset addIs [equalityI,contract.K] addSEs [contract_combE2]) 1);
qed "field_contract_eq";
bind_thm ("reduction_refl",
@@ -60,37 +60,36 @@
val S_contractE = contract.mk_cases comb.con_defs "S -1-> r";
val Ap_contractE = contract.mk_cases comb.con_defs "p#q -1-> r";
-val contract_cs =
- ZF_cs addSIs comb.intrs
- addIs contract.intrs
- addSEs [contract_combD1,contract_combD2] (*type checking*)
- addSEs [K_contractE, S_contractE, Ap_contractE]
- addSEs comb.free_SEs;
+AddSIs comb.intrs;
+AddIs contract.intrs;
+AddSEs [contract_combD1,contract_combD2]; (*type checking*)
+AddSEs [K_contractE, S_contractE, Ap_contractE];
+AddSEs comb.free_SEs;
goalw Comb.thy [I_def] "!!r. I -1-> r ==> P";
-by (fast_tac contract_cs 1);
+by (Fast_tac 1);
qed "I_contract_E";
goal Comb.thy "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)";
-by (fast_tac contract_cs 1);
+by (Fast_tac 1);
qed "K1_contractD";
goal Comb.thy "!!p r. [| p ---> q; r: comb |] ==> p#r ---> q#r";
by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
by (etac rtrancl_induct 1);
-by (fast_tac (contract_cs addIs reduction_rls) 1);
+by (fast_tac (!claset addIs reduction_rls) 1);
by (etac (trans_rtrancl RS transD) 1);
-by (fast_tac (contract_cs addIs reduction_rls) 1);
+by (fast_tac (!claset addIs reduction_rls) 1);
qed "Ap_reduce1";
goal Comb.thy "!!p r. [| p ---> q; r: comb |] ==> r#p ---> r#q";
by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
by (etac rtrancl_induct 1);
-by (fast_tac (contract_cs addIs reduction_rls) 1);
+by (fast_tac (!claset addIs reduction_rls) 1);
by (etac (trans_rtrancl RS transD) 1);
-by (fast_tac (contract_cs addIs reduction_rls) 1);
+by (fast_tac (!claset addIs reduction_rls) 1);
qed "Ap_reduce2";
(** Counterexample to the diamond property for -1-> **)
@@ -108,7 +107,7 @@
qed "KIII_contract3";
goalw Comb.thy [diamond_def] "~ diamond(contract)";
-by (fast_tac (ZF_cs addIs [KIII_contract1,KIII_contract2,KIII_contract3]
+by (fast_tac (!claset addIs [KIII_contract1,KIII_contract2,KIII_contract3]
addSEs [I_contract_E]) 1);
qed "not_diamond_contract";
@@ -122,7 +121,7 @@
val parcontract_combD2 = parcontract.dom_subset RS subsetD RS SigmaD2;
goal Comb.thy "field(parcontract) = comb";
-by (fast_tac (ZF_cs addIs [equalityI, parcontract.K]
+by (fast_tac (!claset addIs [equalityI, parcontract.K]
addSEs [parcontract_combE2]) 1);
qed "field_parcontract_eq";
@@ -131,27 +130,25 @@
val S_parcontractE = parcontract.mk_cases comb.con_defs "S =1=> r";
val Ap_parcontractE = parcontract.mk_cases comb.con_defs "p#q =1=> r";
-val parcontract_cs =
- ZF_cs addSIs comb.intrs
- addIs parcontract.intrs
- addSEs [Ap_E, K_parcontractE, S_parcontractE,
- Ap_parcontractE]
- addSEs [parcontract_combD1, parcontract_combD2] (*type checking*)
- addSEs comb.free_SEs;
+AddSIs comb.intrs;
+AddIs parcontract.intrs;
+AddSEs [Ap_E, K_parcontractE, S_parcontractE, Ap_parcontractE];
+AddSEs [parcontract_combD1, parcontract_combD2]; (*type checking*)
+AddSEs comb.free_SEs;
(*** Basic properties of parallel contraction ***)
goal Comb.thy "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')";
-by (fast_tac parcontract_cs 1);
+by (Fast_tac 1);
qed "K1_parcontractD";
goal Comb.thy "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')";
-by (fast_tac parcontract_cs 1);
+by (Fast_tac 1);
qed "S1_parcontractD";
goal Comb.thy
"!!p q r. S#p#q =1=> r ==> (EX p' q'. r = S#p'#q' & p =1=> p' & q =1=> q')";
-by (fast_tac (parcontract_cs addSDs [S1_parcontractD]) 1);
+by (fast_tac (!claset addSDs [S1_parcontractD]) 1);
qed "S2_parcontractD";
(*Church-Rosser property for parallel contraction*)
@@ -159,7 +156,7 @@
by (rtac (impI RS allI RS allI) 1);
by (etac parcontract.induct 1);
by (ALLGOALS
- (fast_tac (parcontract_cs addSDs [K1_parcontractD,S2_parcontractD])));
+ (fast_tac (!claset addSDs [K1_parcontractD,S2_parcontractD])));
qed "diamond_parcontract";
(*** Transitive closure preserves the Church-Rosser property ***)
@@ -168,8 +165,8 @@
"!!x y r. [| diamond(r); <x,y>:r^+ |] ==> \
\ ALL y'. <x,y'>:r --> (EX z. <y',z>: r^+ & <y,z>: r)";
by (etac trancl_induct 1);
-by (fast_tac (ZF_cs addIs [r_into_trancl]) 1);
-by (slow_best_tac (ZF_cs addSDs [spec RS mp]
+by (fast_tac (!claset addIs [r_into_trancl]) 1);
+by (slow_best_tac (!claset addSDs [spec RS mp]
addIs [r_into_trancl, trans_trancl RS transD]) 1);
val diamond_strip_lemmaE = result() RS spec RS mp RS exE;
@@ -178,7 +175,7 @@
by (rtac (impI RS allI RS allI) 1);
by (etac trancl_induct 1);
by (ALLGOALS
- (slow_best_tac (ZF_cs addIs [r_into_trancl, trans_trancl RS transD]
+ (slow_best_tac (!claset addIs [r_into_trancl, trans_trancl RS transD]
addEs [major RS diamond_strip_lemmaE])));
qed "diamond_trancl";
@@ -187,28 +184,28 @@
goal Comb.thy "!!p q. p-1->q ==> p=1=>q";
by (etac contract.induct 1);
-by (ALLGOALS (fast_tac (parcontract_cs)));
+by (ALLGOALS (fast_tac (!claset)));
qed "contract_imp_parcontract";
goal Comb.thy "!!p q. p--->q ==> p===>q";
by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
by (etac rtrancl_induct 1);
-by (fast_tac (parcontract_cs addIs [r_into_trancl]) 1);
-by (fast_tac (ZF_cs addIs [contract_imp_parcontract,
+by (fast_tac (!claset addIs [r_into_trancl]) 1);
+by (fast_tac (!claset addIs [contract_imp_parcontract,
r_into_trancl, trans_trancl RS transD]) 1);
qed "reduce_imp_parreduce";
goal Comb.thy "!!p q. p=1=>q ==> p--->q";
by (etac parcontract.induct 1);
-by (fast_tac (contract_cs addIs reduction_rls) 1);
-by (fast_tac (contract_cs addIs reduction_rls) 1);
-by (fast_tac (contract_cs addIs reduction_rls) 1);
+by (fast_tac (!claset addIs reduction_rls) 1);
+by (fast_tac (!claset addIs reduction_rls) 1);
+by (fast_tac (!claset addIs reduction_rls) 1);
by (rtac (trans_rtrancl RS transD) 1);
by (ALLGOALS
(fast_tac
- (contract_cs addIs [Ap_reduce1, Ap_reduce2]
+ (!claset addIs [Ap_reduce1, Ap_reduce2]
addSEs [parcontract_combD1,parcontract_combD2])));
qed "parcontract_imp_reduce";
--- a/src/ZF/ex/Data.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Data.ML Fri Jan 03 15:01:55 1997 +0100
@@ -27,7 +27,7 @@
goalw Data.thy (data.defs@data.con_defs) "data(univ(A),univ(A)) <= univ(A)";
by (rtac lfp_lowerbound 1);
by (rtac ([A_subset_univ, Un_upper1] MRS subset_trans RS univ_mono) 2);
-by (fast_tac (ZF_cs addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
+by (fast_tac (!claset addSIs [zero_in_univ, Inl_in_univ, Inr_in_univ,
Pair_in_univ]) 1);
qed "data_univ";
--- a/src/ZF/ex/Enum.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Enum.ML Fri Jan 03 15:01:55 1997 +0100
@@ -11,6 +11,6 @@
open Enum;
goal Enum.thy "C00 ~= C01";
-by (simp_tac (ZF_ss addsimps enum.free_iffs) 1);
+by (simp_tac (!simpset addsimps enum.free_iffs) 1);
result();
--- a/src/ZF/ex/Integ.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Integ.ML Fri Jan 03 15:01:55 1997 +0100
@@ -34,13 +34,13 @@
goalw Integ.thy [intrel_def]
"<<x1,y1>,<x2,y2>>: intrel <-> \
\ x1: nat & y1: nat & x2: nat & y2: nat & x1#+y2 = x2#+y1";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "intrel_iff";
goalw Integ.thy [intrel_def]
"!!x1 x2. [| x1#+y2 = x2#+y1; x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
\ <<x1,y1>,<x2,y2>>: intrel";
-by (fast_tac (ZF_cs addIs prems) 1);
+by (fast_tac (!claset addIs prems) 1);
qed "intrelI";
(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*)
@@ -48,7 +48,7 @@
"p: intrel --> (EX x1 y1 x2 y2. \
\ p = <<x1,y1>,<x2,y2>> & x1#+y2 = x2#+y1 & \
\ x1: nat & y1: nat & x2: nat & y2: nat)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "intrelE_lemma";
val [major,minor] = goal Integ.thy
@@ -60,18 +60,18 @@
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
qed "intrelE";
-val intrel_cs = ZF_cs addSIs [intrelI] addSEs [intrelE];
+AddSIs [intrelI];
+AddSEs [intrelE];
goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def]
"equiv(nat*nat, intrel)";
-by (fast_tac (intrel_cs addSEs [sym, integ_trans_lemma]) 1);
+by (fast_tac (!claset addSEs [sym, integ_trans_lemma]) 1);
qed "equiv_intrel";
-val intrel_ss =
- arith_ss addsimps [equiv_intrel RS eq_equiv_class_iff, intrel_iff,
- add_0_right, add_succ_right]
- addcongs [conj_cong];
+Addsimps [equiv_intrel RS eq_equiv_class_iff, intrel_iff,
+ add_0_right, add_succ_right];
+Addcongs [conj_cong];
val eq_intrelD = equiv_intrel RSN (2,eq_equiv_class);
@@ -79,14 +79,14 @@
goalw Integ.thy [integ_def,quotient_def,znat_def]
"!!m. m : nat ==> $#m : integ";
-by (fast_tac (ZF_cs addSIs [nat_0I]) 1);
+by (fast_tac (!claset addSIs [nat_0I]) 1);
qed "znat_type";
goalw Integ.thy [znat_def]
"!!m n. [| $#m = $#n; n: nat |] ==> m=n";
by (dtac eq_intrelD 1);
by (typechk_tac [nat_0I, SigmaI]);
-by (asm_full_simp_tac intrel_ss 1);
+by (Asm_full_simp_tac 1);
qed "znat_inject";
@@ -94,8 +94,8 @@
goalw Integ.thy [congruent_def]
"congruent(intrel, %<x,y>. intrel``{<y,x>})";
-by (safe_tac intrel_cs);
-by (asm_full_simp_tac (intrel_ss addsimps add_ac) 1);
+by (safe_tac (!claset));
+by (asm_full_simp_tac (!simpset addsimps add_ac) 1);
qed "zminus_congruent";
(*Resolve th against the corresponding facts for zminus*)
@@ -110,24 +110,24 @@
goalw Integ.thy [integ_def,zminus_def]
"!!z w. [| $~z = $~w; z: integ; w: integ |] ==> z=w";
by (etac (zminus_ize UN_equiv_class_inject) 1);
-by (safe_tac intrel_cs);
+by (safe_tac (!claset));
(*The setloop is only needed because assumptions are in the wrong order!*)
-by (asm_full_simp_tac (intrel_ss addsimps add_ac
+by (asm_full_simp_tac (!simpset addsimps add_ac
setloop dtac eq_intrelD) 1);
qed "zminus_inject";
goalw Integ.thy [zminus_def]
"!!x y.[| x: nat; y: nat |] ==> $~ (intrel``{<x,y>}) = intrel `` {<y,x>}";
-by (asm_simp_tac (ZF_ss addsimps [zminus_ize UN_equiv_class, SigmaI]) 1);
+by (asm_simp_tac (!simpset addsimps [zminus_ize UN_equiv_class, SigmaI]) 1);
qed "zminus";
goalw Integ.thy [integ_def] "!!z. z : integ ==> $~ ($~ z) = z";
by (REPEAT (eresolve_tac [quotientE,SigmaE,ssubst] 1));
-by (asm_simp_tac (ZF_ss addsimps [zminus]) 1);
+by (asm_simp_tac (!simpset addsimps [zminus]) 1);
qed "zminus_zminus";
goalw Integ.thy [integ_def, znat_def] "$~ ($#0) = $#0";
-by (simp_tac (arith_ss addsimps [zminus]) 1);
+by (simp_tac (!simpset addsimps [zminus]) 1);
qed "zminus_0";
@@ -135,16 +135,16 @@
(*No natural number is negative!*)
goalw Integ.thy [znegative_def, znat_def] "~ znegative($# n)";
-by (safe_tac intrel_cs);
+by (safe_tac (!claset));
by (dres_inst_tac [("psi", "?lhs=?rhs")] asm_rl 1);
by (dres_inst_tac [("psi", "?lhs<?rhs")] asm_rl 1);
-by (fast_tac (FOL_cs addss
- (intrel_ss addsimps [add_le_self2 RS le_imp_not_lt])) 1);
+by (fast_tac (!claset addss
+ (!simpset addsimps [add_le_self2 RS le_imp_not_lt])) 1);
qed "not_znegative_znat";
goalw Integ.thy [znegative_def, znat_def]
"!!n. n: nat ==> znegative($~ $# succ(n))";
-by (asm_simp_tac (intrel_ss addsimps [zminus]) 1);
+by (asm_simp_tac (!simpset addsimps [zminus]) 1);
by (REPEAT
(ares_tac [refl, exI, conjI, nat_0_le,
refl RS intrelI RS imageI, consI1, nat_0I, nat_succI] 1));
@@ -155,20 +155,21 @@
goalw Integ.thy [congruent_def]
"congruent(intrel, %<x,y>. (y#-x) #+ (x#-y))";
-by (safe_tac intrel_cs);
-by (ALLGOALS (asm_simp_tac intrel_ss));
+by (safe_tac (!claset));
+by (ALLGOALS Asm_simp_tac);
by (etac rev_mp 1);
by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1 THEN
REPEAT (assume_tac 1));
-by (asm_simp_tac intrel_ss 3);
+by (Asm_simp_tac 3);
by (asm_simp_tac (*this one's very sensitive to order of rewrites*)
- (arith_ss0 addsimps [diff_add_inverse,diff_add_0,add_0_right]) 2);
-by (asm_simp_tac intrel_ss 1);
+ (!simpset delsimps [add_succ_right]
+ addsimps [diff_add_inverse,diff_add_0]) 2);
+by (Asm_simp_tac 1);
by (rtac impI 1);
by (etac subst 1);
by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1 THEN
REPEAT (assume_tac 1));
-by (asm_simp_tac (arith_ss addsimps [diff_add_inverse, diff_add_0]) 1);
+by (asm_simp_tac (!simpset addsimps [diff_add_inverse, diff_add_0]) 1);
qed "zmagnitude_congruent";
@@ -184,17 +185,18 @@
goalw Integ.thy [zmagnitude_def]
"!!x y. [| x: nat; y: nat |] ==> \
\ zmagnitude (intrel``{<x,y>}) = (y #- x) #+ (x #- y)";
-by (asm_simp_tac (ZF_ss addsimps [zmagnitude_ize UN_equiv_class, SigmaI]) 1);
+by (asm_simp_tac
+ (!simpset addsimps [zmagnitude_ize UN_equiv_class, SigmaI]) 1);
qed "zmagnitude";
goalw Integ.thy [znat_def]
"!!n. n: nat ==> zmagnitude($# n) = n";
-by (asm_simp_tac (intrel_ss addsimps [zmagnitude]) 1);
+by (asm_simp_tac (!simpset addsimps [zmagnitude]) 1);
qed "zmagnitude_znat";
goalw Integ.thy [znat_def]
"!!n. n: nat ==> zmagnitude($~ $# n) = n";
-by (asm_simp_tac (intrel_ss addsimps [zmagnitude, zminus]) 1);
+by (asm_simp_tac (!simpset addsimps [zmagnitude, zminus]) 1);
qed "zmagnitude_zminus_znat";
@@ -207,14 +209,14 @@
\ let <x1,y1>=z1; <x2,y2>=z2 \
\ in intrel``{<x1#+x2, y1#+y2>})";
(*Proof via congruent2_commuteI seems longer*)
-by (safe_tac intrel_cs);
-by (asm_simp_tac (intrel_ss addsimps [add_assoc, Let_def]) 1);
+by (safe_tac (!claset));
+by (asm_simp_tac (!simpset addsimps [add_assoc, Let_def]) 1);
(*The rest should be trivial, but rearranging terms is hard;
add_ac does not help rewriting with the assumptions.*)
by (res_inst_tac [("m1","x1a")] (add_left_commute RS ssubst) 1);
by (res_inst_tac [("m1","x2a")] (add_left_commute RS ssubst) 3);
by (typechk_tac [add_type]);
-by (asm_simp_tac (arith_ss addsimps [add_assoc RS sym]) 1);
+by (asm_simp_tac (!simpset addsimps [add_assoc RS sym]) 1);
qed "zadd_congruent2";
(*Resolve th against the corresponding facts for zadd*)
@@ -223,7 +225,7 @@
goalw Integ.thy [integ_def,zadd_def]
"!!z w. [| z: integ; w: integ |] ==> z $+ w : integ";
by (rtac (zadd_ize UN_equiv_class_type2) 1);
-by (simp_tac (ZF_ss addsimps [Let_def]) 3);
+by (simp_tac (!simpset addsimps [Let_def]) 3);
by (REPEAT (ares_tac [split_type, add_type, quotientI, SigmaI] 1));
qed "zadd_type";
@@ -231,25 +233,25 @@
"!!x1 y1. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
\ (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) = \
\ intrel `` {<x1#+x2, y1#+y2>}";
-by (asm_simp_tac (ZF_ss addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1);
-by (simp_tac (ZF_ss addsimps [Let_def]) 1);
+by (asm_simp_tac (!simpset addsimps [zadd_ize UN_equiv_class2, SigmaI]) 1);
+by (simp_tac (!simpset addsimps [Let_def]) 1);
qed "zadd";
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $+ z = z";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (arith_ss addsimps [zadd]) 1);
+by (asm_simp_tac (!simpset addsimps [zadd]) 1);
qed "zadd_0";
goalw Integ.thy [integ_def]
"!!z w. [| z: integ; w: integ |] ==> $~ (z $+ w) = $~ z $+ $~ w";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (arith_ss addsimps [zminus,zadd]) 1);
+by (asm_simp_tac (!simpset addsimps [zminus,zadd]) 1);
qed "zminus_zadd_distrib";
goalw Integ.thy [integ_def]
"!!z w. [| z: integ; w: integ |] ==> z $+ w = w $+ z";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (intrel_ss addsimps (add_ac @ [zadd])) 1);
+by (asm_simp_tac (!simpset addsimps (add_ac @ [zadd])) 1);
qed "zadd_commute";
goalw Integ.thy [integ_def]
@@ -257,31 +259,31 @@
\ (z1 $+ z2) $+ z3 = z1 $+ (z2 $+ z3)";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
(*rewriting is much faster without intrel_iff, etc.*)
-by (asm_simp_tac (arith_ss addsimps [zadd, add_assoc]) 1);
+by (asm_simp_tac (!simpset addsimps [zadd, add_assoc]) 1);
qed "zadd_assoc";
(*For AC rewriting*)
qed_goal "zadd_left_commute" Integ.thy
"!!z1 z2 z3. [| z1:integ; z2:integ; z3: integ |] ==> \
\ z1$+(z2$+z3) = z2$+(z1$+z3)"
- (fn _ => [asm_simp_tac (ZF_ss addsimps [zadd_assoc RS sym, zadd_commute]) 1]);
+ (fn _ => [asm_simp_tac (!simpset addsimps [zadd_assoc RS sym, zadd_commute]) 1]);
(*Integer addition is an AC operator*)
val zadd_ac = [zadd_assoc, zadd_commute, zadd_left_commute];
goalw Integ.thy [znat_def]
"!!m n. [| m: nat; n: nat |] ==> $# (m #+ n) = ($#m) $+ ($#n)";
-by (asm_simp_tac (arith_ss addsimps [zadd]) 1);
+by (asm_simp_tac (!simpset addsimps [zadd]) 1);
qed "znat_add";
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> z $+ ($~ z) = $#0";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (intrel_ss addsimps [zminus, zadd, add_commute]) 1);
+by (asm_simp_tac (!simpset addsimps [zminus, zadd, add_commute]) 1);
qed "zadd_zminus_inverse";
goal Integ.thy "!!z. z : integ ==> ($~ z) $+ z = $#0";
by (asm_simp_tac
- (ZF_ss addsimps [zadd_commute, zminus_type, zadd_zminus_inverse]) 1);
+ (!simpset addsimps [zadd_commute, zminus_type, zadd_zminus_inverse]) 1);
qed "zadd_zminus_inverse2";
goal Integ.thy "!!z. z:integ ==> z $+ $#0 = z";
@@ -303,15 +305,15 @@
\ split(%x1 y1. split(%x2 y2. \
\ intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))";
by (rtac (equiv_intrel RS congruent2_commuteI) 1);
-by (safe_tac intrel_cs);
-by (ALLGOALS (asm_simp_tac intrel_ss));
+by (safe_tac (!claset));
+by (ALLGOALS Asm_simp_tac);
(*Proof that zmult is congruent in one argument*)
by (asm_simp_tac
- (arith_ss addsimps (add_ac @ [add_mult_distrib_left RS sym])) 2);
+ (!simpset addsimps (add_ac @ [add_mult_distrib_left RS sym])) 2);
by (asm_simp_tac
- (arith_ss addsimps ([add_assoc RS sym, add_mult_distrib_left RS sym])) 2);
+ (!simpset addsimps ([add_assoc RS sym, add_mult_distrib_left RS sym])) 2);
(*Proof that zmult is commutative on representatives*)
-by (asm_simp_tac (arith_ss addsimps (mult_ac@add_ac)) 1);
+by (asm_simp_tac (!simpset addsimps (mult_ac@add_ac)) 1);
qed "zmult_congruent2";
@@ -329,36 +331,36 @@
"!!x1 x2. [| x1: nat; y1: nat; x2: nat; y2: nat |] ==> \
\ (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) = \
\ intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}";
-by (asm_simp_tac (ZF_ss addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1);
+by (asm_simp_tac (!simpset addsimps [zmult_ize UN_equiv_class2, SigmaI]) 1);
qed "zmult";
goalw Integ.thy [integ_def,znat_def] "!!z. z : integ ==> $#0 $* z = $#0";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (arith_ss addsimps [zmult]) 1);
+by (asm_simp_tac (!simpset addsimps [zmult]) 1);
qed "zmult_0";
goalw Integ.thy [integ_def,znat_def]
"!!z. z : integ ==> $#1 $* z = z";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (arith_ss addsimps [zmult, add_0_right]) 1);
+by (asm_simp_tac (!simpset addsimps [zmult, add_0_right]) 1);
qed "zmult_1";
goalw Integ.thy [integ_def]
"!!z w. [| z: integ; w: integ |] ==> ($~ z) $* w = $~ (z $* w)";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (intrel_ss addsimps ([zminus, zmult] @ add_ac)) 1);
+by (asm_simp_tac (!simpset addsimps ([zminus, zmult] @ add_ac)) 1);
qed "zmult_zminus";
goalw Integ.thy [integ_def]
"!!z w. [| z: integ; w: integ |] ==> ($~ z) $* ($~ w) = (z $* w)";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (intrel_ss addsimps ([zminus, zmult] @ add_ac)) 1);
+by (asm_simp_tac (!simpset addsimps ([zminus, zmult] @ add_ac)) 1);
qed "zmult_zminus_zminus";
goalw Integ.thy [integ_def]
"!!z w. [| z: integ; w: integ |] ==> z $* w = w $* z";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (intrel_ss addsimps ([zmult] @ add_ac @ mult_ac)) 1);
+by (asm_simp_tac (!simpset addsimps ([zmult] @ add_ac @ mult_ac)) 1);
qed "zmult_commute";
goalw Integ.thy [integ_def]
@@ -366,7 +368,7 @@
\ (z1 $* z2) $* z3 = z1 $* (z2 $* z3)";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
by (asm_simp_tac
- (intrel_ss addsimps ([zmult, add_mult_distrib_left,
+ (!simpset addsimps ([zmult, add_mult_distrib_left,
add_mult_distrib] @ add_ac @ mult_ac)) 1);
qed "zmult_assoc";
@@ -374,7 +376,7 @@
qed_goal "zmult_left_commute" Integ.thy
"!!z1 z2 z3. [| z1:integ; z2:integ; z3: integ |] ==> \
\ z1$*(z2$*z3) = z2$*(z1$*z3)"
- (fn _ => [asm_simp_tac (ZF_ss addsimps [zmult_assoc RS sym,
+ (fn _ => [asm_simp_tac (!simpset addsimps [zmult_assoc RS sym,
zmult_commute]) 1]);
(*Integer multiplication is an AC operator*)
@@ -384,21 +386,19 @@
"!!z1 z2 z3. [| z1: integ; z2: integ; w: integ |] ==> \
\ (z1 $+ z2) $* w = (z1 $* w) $+ (z2 $* w)";
by (REPEAT (eresolve_tac [quotientE, SigmaE, ssubst] 1));
-by (asm_simp_tac (intrel_ss addsimps [zadd, zmult, add_mult_distrib]) 1);
-by (asm_simp_tac (intrel_ss addsimps (add_ac @ mult_ac)) 1);
+by (asm_simp_tac (!simpset addsimps [zadd, zmult, add_mult_distrib]) 1);
+by (asm_simp_tac (!simpset addsimps (add_ac @ mult_ac)) 1);
qed "zadd_zmult_distrib";
val integ_typechecks =
[znat_type, zminus_type, zmagnitude_type, zadd_type, zmult_type];
-val zadd_simps =
- [zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2];
+Addsimps integ_typechecks;
-val zminus_simps = [zminus_zminus, zminus_0];
+Addsimps [zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2];
-val zmult_simps = [zmult_0, zmult_1, zmult_zminus];
+Addsimps [zminus_zminus, zminus_0];
-val integ_ss =
- arith_ss addsimps (zadd_simps @ zminus_simps @ zmult_simps @
- [zmagnitude_znat, zmagnitude_zminus_znat] @
- integ_typechecks);
+Addsimps [zmult_0, zmult_1, zmult_zminus];
+
+Addsimps [zmagnitude_znat, zmagnitude_zminus_znat];
--- a/src/ZF/ex/LList.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/LList.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,6 +8,12 @@
open LList;
+Delrules [subsetI, subsetCE];
+AddSIs [subset_refl, cons_subsetI, subset_consI,
+ Union_least, UN_least, Un_least,
+ Inter_greatest, Int_greatest, RepFun_subset,
+ Un_upper1, Un_upper2, Int_lower1, Int_lower2];
+
(*An elimination rule, for type-checking*)
val LConsE = llist.mk_cases llist.con_defs "LCons(a,l) : llist(A)";
@@ -17,7 +23,7 @@
goal LList.thy "llist(A) = {0} <+> (A <*> llist(A))";
let open llist; val rew = rewrite_rule con_defs in
-by (fast_tac (qsum_cs addSIs (equalityI :: map rew intrs)
+by (fast_tac (!claset addSIs (subsetI :: equalityI :: map rew intrs)
addEs [rew elim]) 1)
end;
qed "llist_unfold";
@@ -32,12 +38,11 @@
(** Closure of quniv(A) under llist -- why so complex? Its a gfp... **)
-val quniv_cs = subset_cs addSIs [QPair_Int_Vset_subset_UN RS subset_trans,
+AddSIs [QPair_Int_Vset_subset_UN RS subset_trans,
QPair_subset_univ,
- empty_subsetI, one_in_quniv RS qunivD]
- addIs [Int_lower1 RS subset_trans]
- addSDs [qunivD]
- addSEs [Ord_in_Ord];
+ empty_subsetI, one_in_quniv RS qunivD];
+AddSDs [qunivD];
+AddSEs [Ord_in_Ord];
goal LList.thy
"!!i. Ord(i) ==> ALL l: llist(quniv(A)). l Int Vset(i) <= univ(eclose(A))";
@@ -46,10 +51,9 @@
by (etac llist.elim 1);
by (rewrite_goals_tac ([QInl_def,QInr_def]@llist.con_defs));
(*LNil case*)
-by (fast_tac quniv_cs 1);
+by (Asm_simp_tac 1);
(*LCons case*)
-by (safe_tac quniv_cs);
-by (ALLGOALS (fast_tac (quniv_cs addSEs [Ord_trans, make_elim bspec])));
+by (deepen_tac (!claset addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
qed "llist_quniv_lemma";
goal LList.thy "llist(quniv(A)) <= quniv(A)";
@@ -64,9 +68,8 @@
(*** Lazy List Equality: lleq ***)
-val lleq_cs = subset_cs
- addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono]
- addSEs [Ord_in_Ord, Pair_inject];
+AddSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono];
+AddSEs [Ord_in_Ord, Pair_inject];
(*Lemma for proving finality. Unfold the lazy list; use induction hypothesis*)
goal LList.thy
@@ -75,7 +78,7 @@
by (REPEAT (resolve_tac [allI, impI] 1));
by (etac lleq.elim 1);
by (rewrite_goals_tac (QInr_def::llist.con_defs));
-by (safe_tac lleq_cs);
+by (safe_tac (!claset));
by (fast_tac (subset_cs addSEs [Ord_trans, make_elim bspec]) 1);
qed "lleq_Int_Vset_subset_lemma";
@@ -87,9 +90,9 @@
val [prem] = goal LList.thy "<l,l'> : lleq(A) ==> <l',l> : lleq(A)";
by (rtac (prem RS converseI RS lleq.coinduct) 1);
by (rtac (lleq.dom_subset RS converse_type) 1);
-by (safe_tac converse_cs);
+by (safe_tac (!claset));
by (etac lleq.elim 1);
-by (ALLGOALS (fast_tac qconverse_cs));
+by (ALLGOALS Fast_tac);
qed "lleq_symmetric";
goal LList.thy "!!l l'. <l,l'> : lleq(A) ==> l=l'";
@@ -102,9 +105,9 @@
"[| l=l'; l: llist(A) |] ==> <l,l'> : lleq(A)";
by (res_inst_tac [("X", "{<l,l>. l: llist(A)}")] lleq.coinduct 1);
by (rtac (lprem RS RepFunI RS (eqprem RS subst)) 1);
-by (safe_tac qpair_cs);
+by (safe_tac (!claset));
by (etac llist.elim 1);
-by (ALLGOALS (fast_tac pair_cs));
+by (ALLGOALS Fast_tac);
qed "equal_llist_implies_leq";
@@ -137,19 +140,19 @@
goal LList.thy "!!a A. a:A ==> lconst(a): llist(A)";
by (rtac (singletonI RS llist.coinduct) 1);
by (etac (lconst_in_quniv RS singleton_subsetI) 1);
-by (fast_tac (ZF_cs addSIs [lconst]) 1);
+by (fast_tac (!claset addSIs [lconst]) 1);
qed "lconst_type";
(*** flip --- equations merely assumed; certain consequences proved ***)
-val flip_ss = ZF_ss addsimps [flip_LNil, flip_LCons, not_type];
+Addsimps [flip_LNil, flip_LCons, not_type];
goal QUniv.thy "!!b. b:bool ==> b Int X <= univ(eclose(A))";
-by (fast_tac (quniv_cs addSEs [boolE]) 1);
+by (fast_tac (!claset addIs [Int_lower1 RS subset_trans] addSEs [boolE]) 1);
qed "bool_Int_subset_univ";
-val flip_cs = quniv_cs addSIs [not_type]
- addIs [bool_Int_subset_univ];
+AddSIs [not_type];
+AddIs [bool_Int_subset_univ];
(*Reasoning borrowed from lleq.ML; a similar proof works for all
"productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
@@ -159,14 +162,11 @@
by (etac trans_induct 1);
by (rtac ballI 1);
by (etac llist.elim 1);
-by (asm_simp_tac flip_ss 1);
-by (asm_simp_tac flip_ss 2);
-by (rewrite_goals_tac ([QInl_def,QInr_def]@llist.con_defs));
-(*LNil case*)
-by (fast_tac flip_cs 1);
+by (ALLGOALS Asm_simp_tac);
+by (ALLGOALS
+ (asm_simp_tac (!simpset addsimps ([QInl_def,QInr_def]@llist.con_defs))));
(*LCons case*)
-by (safe_tac flip_cs);
-by (ALLGOALS (fast_tac (flip_cs addSEs [Ord_trans, make_elim bspec])));
+by (deepen_tac (!claset addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
qed "flip_llist_quniv_lemma";
goal LList.thy "!!l. l: llist(bool) ==> flip(l) : quniv(bool)";
@@ -178,12 +178,11 @@
by (res_inst_tac [("X", "{flip(l) . l:llist(bool)}")]
llist.coinduct 1);
by (rtac (prem RS RepFunI) 1);
-by (fast_tac (ZF_cs addSIs [flip_in_quniv]) 1);
+by (fast_tac (!claset addSIs [flip_in_quniv]) 1);
by (etac RepFunE 1);
by (etac llist.elim 1);
-by (asm_simp_tac flip_ss 1);
-by (asm_simp_tac flip_ss 1);
-by (fast_tac (ZF_cs addSIs [not_type]) 1);
+by (ALLGOALS Asm_simp_tac);
+by (Fast_tac 1);
qed "flip_type";
val [prem] = goal LList.thy
@@ -191,10 +190,10 @@
by (res_inst_tac [("X1", "{<flip(flip(l)),l> . l:llist(bool)}")]
(lleq.coinduct RS lleq_implies_equal) 1);
by (rtac (prem RS RepFunI) 1);
-by (fast_tac (ZF_cs addSIs [flip_type]) 1);
+by (fast_tac (!claset addSIs [flip_type]) 1);
by (etac RepFunE 1);
by (etac llist.elim 1);
-by (asm_simp_tac flip_ss 1);
-by (asm_simp_tac (flip_ss addsimps [flip_type, not_not]) 1);
-by (fast_tac (ZF_cs addSIs [not_type]) 1);
+by (Asm_simp_tac 1);
+by (asm_simp_tac (!simpset addsimps [flip_type, not_not]) 1);
+by (fast_tac (!claset addSIs [not_type]) 1);
qed "flip_flip";
--- a/src/ZF/ex/Limit.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Limit.ML Fri Jan 03 15:01:55 1997 +0100
@@ -5,6 +5,8 @@
The inverse limit construction.
*)
+val nat_linear_le = [nat_into_Ord,nat_into_Ord] MRS Ord_linear_le;
+
open Limit;
(*----------------------------------------------------------------------*)
@@ -16,16 +18,6 @@
fun rotate n i = EVERY(replicate n (etac revcut_rl i));
(*----------------------------------------------------------------------*)
-(* Preliminary theorems. *)
-(*----------------------------------------------------------------------*)
-
-val theI2 = prove_goal ZF.thy (* From Larry *)
- "[| EX! x. P(x); !!x. P(x) ==> Q(x) |] ==> Q(THE x.P(x))"
- (fn prems => [ resolve_tac prems 1,
- rtac theI 1,
- resolve_tac prems 1 ]);
-
-(*----------------------------------------------------------------------*)
(* Basic results. *)
(*----------------------------------------------------------------------*)
@@ -78,13 +70,13 @@
\ rel(D,x,z); \
\ !!x y. [| rel(D,x,y); rel(D,y,x); x:set(D); y:set(D)|] ==> x=y |] ==> \
\ po(D)";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
brr prems 1;
val poI = result();
val prems = goalw Limit.thy [cpo_def]
"[| po(D); !!X. chain(D,X) ==> islub(D,X,x(D,X))|] ==> cpo(D)";
-by (safe_tac (lemmas_cs addSIs [exI]));
+by (safe_tac (!claset addSIs [exI]));
brr prems 1;
val cpoI = result();
@@ -112,7 +104,7 @@
val cpo_antisym = result();
val [cpo,chain,ex] = goalw Limit.thy [cpo_def] (* cpo_islub *)
- "[|cpo(D); chain(D,X);!!x. islub(D,X,x) ==> R|] ==> R";
+ "[|cpo(D); chain(D,X); !!x. islub(D,X,x) ==> R|] ==> R";
by (rtac (chain RS (cpo RS conjunct2 RS spec RS mp) RS exE) 1);
brr[ex]1; (* above theorem would loop *)
val cpo_islub = result();
@@ -123,7 +115,7 @@
val prems = goalw Limit.thy [islub_def] (* islub_isub *)
"islub(D,X,x) ==> isub(D,X,x)";
-by (simp_tac (ZF_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
val islub_isub = result();
val prems = goal Limit.thy
@@ -146,30 +138,30 @@
val prems = goalw Limit.thy [islub_def] (* islubI *)
"[|isub(D,X,x); !!y. isub(D,X,y) ==> rel(D,x,y)|] ==> islub(D,X,x)";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (REPEAT(ares_tac prems 1));
val islubI = result();
val prems = goalw Limit.thy [isub_def] (* isubI *)
- "[|x:set(D);!!n. n:nat ==> rel(D,X`n,x)|] ==> isub(D,X,x)";
-by (safe_tac lemmas_cs);
+ "[|x:set(D); !!n. n:nat ==> rel(D,X`n,x)|] ==> isub(D,X,x)";
+by (safe_tac (!claset));
by (REPEAT(ares_tac prems 1));
val isubI = result();
val prems = goalw Limit.thy [isub_def] (* isubE *)
- "!!z.[|isub(D,X,x);[|x:set(D);!!n.n:nat==>rel(D,X`n,x)|] ==> P|] ==> P";
-by (safe_tac lemmas_cs);
-by (asm_simp_tac ZF_ss 1);
+ "!!z.[|isub(D,X,x);[|x:set(D); !!n.n:nat==>rel(D,X`n,x)|] ==> P|] ==> P";
+by (safe_tac (!claset));
+by (Asm_simp_tac 1);
val isubE = result();
val prems = goalw Limit.thy [isub_def] (* isubD1 *)
"isub(D,X,x) ==> x:set(D)";
-by (simp_tac (ZF_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
val isubD1 = result();
val prems = goalw Limit.thy [isub_def] (* isubD2 *)
"[|isub(D,X,x); n:nat|]==>rel(D,X`n,x)";
-by (simp_tac (ZF_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
val isubD2 = result();
val prems = goal Limit.thy
@@ -200,12 +192,12 @@
(*----------------------------------------------------------------------*)
val chainI = prove_goalw Limit.thy [chain_def]
- "!!z.[|X:nat->set(D);!!n. n:nat ==> rel(D,X`n,X`succ(n))|] ==> chain(D,X)"
- (fn prems => [asm_simp_tac ZF_ss 1]);
+ "!!z.[|X:nat->set(D); !!n. n:nat ==> rel(D,X`n,X`succ(n))|] ==> chain(D,X)"
+ (fn prems => [Asm_simp_tac 1]);
val prems = goalw Limit.thy [chain_def]
"chain(D,X) ==> X : nat -> set(D)";
-by (asm_simp_tac (ZF_ss addsimps prems) 1);
+by (asm_simp_tac (!simpset addsimps prems) 1);
val chain_fun = result();
val prems = goalw Limit.thy [chain_def]
@@ -223,7 +215,7 @@
val prems = goal Limit.thy (* chain_rel_gen_add *)
"[|chain(D,X); cpo(D); n:nat; m:nat|] ==> rel(D,X`n,(X`(m #+ n)))";
by (res_inst_tac [("n","m")] nat_induct 1);
-by (ALLGOALS(simp_tac arith_ss));
+by (ALLGOALS Simp_tac);
by (rtac cpo_trans 3); (* loops if repeated *)
brr(cpo_refl::chain_in::chain_rel::nat_succI::add_type::prems) 1;
val chain_rel_gen_add = result();
@@ -232,7 +224,7 @@
"[| n le succ(x); ~ n le x; x : nat; n:nat |] ==> n = succ(x)";
by (rtac le_anti_sym 1);
by (resolve_tac prems 1);
-by (simp_tac arith_ss 1);
+by (Simp_tac 1);
by (rtac (not_le_iff_lt RS iffD1) 1);
by (REPEAT(resolve_tac (nat_into_Ord::prems) 1));
val le_succ_eq = result();
@@ -243,8 +235,8 @@
by (assume_tac 3);
by (rtac (hd prems) 2);
by (res_inst_tac [("n","m")] nat_induct 1);
-by (safe_tac lemmas_cs);
-by (asm_full_simp_tac (arith_ss addsimps prems) 2);
+by (safe_tac (!claset));
+by (asm_full_simp_tac (!simpset addsimps prems) 2);
by (rtac cpo_trans 4);
by (rtac (le_succ_eq RS subst) 3);
brr(cpo_refl::chain_in::chain_rel::nat_0I::nat_succI::prems) 1;
@@ -271,7 +263,7 @@
"pcpo(D) ==> EX! x. x:set(D) & (ALL y:set(D). rel(D,x,y))";
by (rtac (hd prems RS conjunct2 RS bexE) 1);
by (rtac ex1I 1);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (assume_tac 1);
by (etac bspec 1);
by (assume_tac 1);
@@ -318,14 +310,14 @@
by (rtac lam_type 1);
by (resolve_tac prems 1);
by (rtac ballI 1);
-by (asm_simp_tac (ZF_ss addsimps [nat_succI]) 1);
+by (asm_simp_tac (!simpset addsimps [nat_succI]) 1);
brr(cpo_refl::prems) 1;
val chain_const = result();
val prems = goalw Limit.thy [islub_def,isub_def] (* islub_const *)
"[|x:set(D); cpo(D)|] ==> islub(D,(lam n:nat. x),x)";
-by (simp_tac ZF_ss 1);
-by (safe_tac lemmas_cs);
+by (Simp_tac 1);
+by (safe_tac (!claset));
by (etac bspec 3);
brr(cpo_refl::nat_0I::prems) 1;
val islub_const = result();
@@ -346,8 +338,8 @@
val prems = goalw Limit.thy [isub_def,suffix_def] (* isub_suffix *)
"[|chain(D,X); cpo(D); n:nat|] ==> isub(D,suffix(X,n),x) <-> isub(D,X,x)";
-by (simp_tac (ZF_ss addsimps prems) 1);
-by (safe_tac lemmas_cs);
+by (simp_tac (!simpset addsimps prems) 1);
+by (safe_tac (!claset));
by (dtac bspec 2);
by (assume_tac 3); (* to instantiate unknowns properly *)
by (rtac cpo_trans 1);
@@ -359,12 +351,12 @@
val prems = goalw Limit.thy [islub_def] (* islub_suffix *)
"[|chain(D,X); cpo(D); n:nat|] ==> islub(D,suffix(X,n),x) <-> islub(D,X,x)";
-by (asm_simp_tac (FOL_ss addsimps isub_suffix::prems) 1);
+by (asm_simp_tac (!simpset addsimps isub_suffix::prems) 1);
val islub_suffix = result();
val prems = goalw Limit.thy [lub_def] (* lub_suffix *)
"[|chain(D,X); cpo(D); n:nat|] ==> lub(D,suffix(X,n)) = lub(D,X)";
-by (asm_simp_tac (FOL_ss addsimps islub_suffix::prems) 1);
+by (asm_simp_tac (!simpset addsimps islub_suffix::prems) 1);
val lub_suffix = result();
(*----------------------------------------------------------------------*)
@@ -407,7 +399,7 @@
val dominate_islub = result();
val prems = goalw Limit.thy [subchain_def] (* subchainE *)
- "[|subchain(X,Y); n:nat;!!m. [|m:nat; X`n = Y`(n #+ m)|] ==> Q|] ==> Q";
+ "[|subchain(X,Y); n:nat; !!m. [|m:nat; X`n = Y`(n #+ m)|] ==> Q|] ==> Q";
by (rtac (hd prems RS bspec RS bexE) 1);
by (resolve_tac prems 2);
by (assume_tac 3);
@@ -421,10 +413,10 @@
by (rtac (ub RS isubD1) 1);
by (rtac (subch RS subchainE) 1);
by (assume_tac 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (rtac isubD2 1); (* br with Destruction rule ?? *)
by (resolve_tac prems 1);
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
val subchain_isub = result();
val prems = goal Limit.thy (* dominate_islub_eq *)
@@ -447,7 +439,7 @@
val prems = goalw Limit.thy [matrix_def] (* matrix_fun *)
"matrix(D,M) ==> M : nat -> (nat -> set(D))";
-by (simp_tac (ZF_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
val matrix_fun = result();
val prems = goalw Limit.thy [] (* matrix_in_fun *)
@@ -464,17 +456,17 @@
val prems = goalw Limit.thy [matrix_def] (* matrix_rel_1_0 *)
"[|matrix(D,M); n:nat; m:nat|] ==> rel(D,M`n`m,M`succ(n)`m)";
-by (simp_tac (ZF_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
val matrix_rel_1_0 = result();
val prems = goalw Limit.thy [matrix_def] (* matrix_rel_0_1 *)
"[|matrix(D,M); n:nat; m:nat|] ==> rel(D,M`n`m,M`n`succ(m))";
-by (simp_tac (ZF_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
val matrix_rel_0_1 = result();
val prems = goalw Limit.thy [matrix_def] (* matrix_rel_1_1 *)
"[|matrix(D,M); n:nat; m:nat|] ==> rel(D,M`n`m,M`succ(n)`succ(m))";
-by (simp_tac (ZF_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
val matrix_rel_1_1 = result();
val prems = goal Limit.thy (* fun_swap *)
@@ -488,24 +480,24 @@
val prems = goalw Limit.thy [matrix_def] (* matrix_sym_axis *)
"!!z. matrix(D,M) ==> matrix(D,lam m:nat. lam n:nat. M`n`m)";
-by (simp_tac arith_ss 1 THEN safe_tac lemmas_cs THEN
-REPEAT(asm_simp_tac (ZF_ss addsimps [fun_swap]) 1));
+by (Simp_tac 1 THEN safe_tac (!claset) THEN
+REPEAT(asm_simp_tac (!simpset addsimps [fun_swap]) 1));
val matrix_sym_axis = result();
val prems = goalw Limit.thy [chain_def] (* matrix_chain_diag *)
"matrix(D,M) ==> chain(D,lam n:nat. M`n`n)";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac lam_type 1);
by (rtac matrix_in 1);
by (REPEAT(ares_tac prems 1));
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (rtac matrix_rel_1_1 1);
by (REPEAT(ares_tac prems 1));
val matrix_chain_diag = result();
val prems = goalw Limit.thy [chain_def] (* matrix_chain_left *)
"[|matrix(D,M); n:nat|] ==> chain(D,M`n)";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac apply_type 1);
by (rtac matrix_fun 1);
by (REPEAT(ares_tac prems 1));
@@ -515,44 +507,44 @@
val prems = goalw Limit.thy [chain_def] (* matrix_chain_right *)
"[|matrix(D,M); m:nat|] ==> chain(D,lam n:nat. M`n`m)";
-by (safe_tac lemmas_cs);
-by (asm_simp_tac(arith_ss addsimps prems) 2);
+by (safe_tac (!claset));
+by (asm_simp_tac(!simpset addsimps prems) 2);
brr(lam_type::matrix_in::matrix_rel_1_0::prems) 1;
val matrix_chain_right = result();
val prems = goalw Limit.thy [matrix_def] (* matrix_chainI *)
- "[|!!x.x:nat==>chain(D,M`x);!!y.y:nat==>chain(D,lam x:nat. M`x`y); \
+ "[|!!x.x:nat==>chain(D,M`x); !!y.y:nat==>chain(D,lam x:nat. M`x`y); \
\ M:nat->nat->set(D); cpo(D)|] ==> matrix(D,M)";
-by (safe_tac (lemmas_cs addSIs [ballI]));
+by (safe_tac (!claset addSIs [ballI]));
by (cut_inst_tac[("y1","m"),("n","n")](hd(tl prems) RS chain_rel) 2);
-by (asm_full_simp_tac arith_ss 4);
+by (Asm_full_simp_tac 4);
by (rtac cpo_trans 5);
by (cut_inst_tac[("y1","m"),("n","n")](hd(tl prems) RS chain_rel) 6);
-by (asm_full_simp_tac arith_ss 8);
+by (Asm_full_simp_tac 8);
by (TRYALL(rtac (chain_fun RS apply_type)));
brr(chain_rel::nat_succI::prems) 1;
val matrix_chainI = result();
val lemma = prove_goal Limit.thy
"!!z.[|m : nat; rel(D, (lam n:nat. M`n`n)`m, y)|] ==> rel(D,M`m`m, y)"
- (fn prems => [asm_full_simp_tac ZF_ss 1]);
+ (fn prems => [Asm_full_simp_tac 1]);
val lemma2 = prove_goal Limit.thy
"!!z.[|x:nat; m:nat; rel(D,(lam n:nat.M`n`m1)`x,(lam n:nat.M`n`m1)`m)|] ==> \
\ rel(D,M`x`m1,M`m`m1)"
- (fn prems => [asm_full_simp_tac ZF_ss 1]);
+ (fn prems => [Asm_full_simp_tac 1]);
val prems = goalw Limit.thy [] (* isub_lemma *)
"[|isub(D,(lam n:nat. M`n`n),y); matrix(D,M); cpo(D)|] ==> \
\ isub(D,(lam n:nat. lub(D,lam m:nat. M`n`m)),y)";
by (rewtac isub_def);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac isubD1 1);
by (resolve_tac prems 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (cut_inst_tac[("a","n")](hd(tl prems) RS matrix_fun RS apply_type) 1);
by (assume_tac 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (rtac islub_least 1);
by (rtac cpo_lub 1);
by (rtac matrix_chain_left 1);
@@ -560,11 +552,11 @@
by (assume_tac 1);
by (resolve_tac prems 1);
by (rewtac isub_def);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac isubD1 1);
by (resolve_tac prems 1);
by (cut_inst_tac[("P","n le na")]excluded_middle 1);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac cpo_trans 1);
by (resolve_tac prems 1);
by (rtac (not_le_iff_lt RS iffD1 RS leI RS chain_rel_gen) 1);
@@ -584,19 +576,19 @@
val prems = goalw Limit.thy [chain_def] (* matrix_chain_lub *)
"[|matrix(D,M); cpo(D)|] ==> chain(D,lam n:nat.lub(D,lam m:nat.M`n`m))";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac lam_type 1);
by (rtac islub_in 1);
by (rtac cpo_lub 1);
by (resolve_tac prems 2);
-by (asm_simp_tac arith_ss 2);
+by (Asm_simp_tac 2);
by (rtac chainI 1);
by (rtac lam_type 1);
by (REPEAT(ares_tac (matrix_in::prems) 1));
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (rtac matrix_rel_0_1 1);
by (REPEAT(ares_tac prems 1));
-by (asm_simp_tac (arith_ss addsimps
+by (asm_simp_tac (!simpset addsimps
[hd prems RS matrix_chain_left RS chain_fun RS eta]) 1);
by (rtac dominate_islub 1);
by (rtac cpo_lub 3);
@@ -621,8 +613,8 @@
by (rtac ballI 1);
by (rtac bexI 1);
by (assume_tac 2);
-by (asm_simp_tac ZF_ss 1);
-by (asm_simp_tac (arith_ss addsimps
+by (Asm_simp_tac 1);
+by (asm_simp_tac (!simpset addsimps
[hd prems RS matrix_chain_left RS chain_fun RS eta]) 1);
by (rtac islub_ub 1);
by (rtac cpo_lub 1);
@@ -635,20 +627,20 @@
val lemma1 = prove_goalw Limit.thy [lub_def]
"lub(D,(lam n:nat. lub(D,lam m:nat. M`n`m))) = \
\ (THE x. islub(D, (lam n:nat. lub(D,lam m:nat. M`n`m)), x))"
- (fn prems => [fast_tac ZF_cs 1]);
+ (fn prems => [Fast_tac 1]);
val lemma2 = prove_goalw Limit.thy [lub_def]
"lub(D,(lam n:nat. M`n`n)) = \
\ (THE x. islub(D, (lam n:nat. M`n`n), x))"
- (fn prems => [fast_tac ZF_cs 1]);
+ (fn prems => [Fast_tac 1]);
val prems = goalw Limit.thy [] (* lub_matrix_diag *)
"[|matrix(D,M); cpo(D)|] ==> \
\ lub(D,(lam n:nat. lub(D,lam m:nat. M`n`m))) = \
\ lub(D,(lam n:nat. M`n`n))";
-by (simp_tac (arith_ss addsimps [lemma1,lemma2]) 1);
+by (simp_tac (!simpset addsimps [lemma1,lemma2]) 1);
by (rewtac islub_def);
-by (simp_tac (FOL_ss addsimps [hd(tl prems) RS (hd prems RS isub_eq)]) 1);
+by (simp_tac (!simpset addsimps [hd(tl prems) RS (hd prems RS isub_eq)]) 1);
val lub_matrix_diag = result();
val [matrix,cpo] = goalw Limit.thy [] (* lub_matrix_diag_sym *)
@@ -656,7 +648,7 @@
\ lub(D,(lam m:nat. lub(D,lam n:nat. M`n`m))) = \
\ lub(D,(lam n:nat. M`n`n))";
by (cut_facts_tac[cpo RS (matrix RS matrix_sym_axis RS lub_matrix_diag)]1);
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val lub_matrix_diag_sym = result();
(*----------------------------------------------------------------------*)
@@ -667,7 +659,7 @@
"[|f:set(D)->set(E); \
\ !!x y. [|rel(D,x,y); x:set(D); y:set(D)|] ==> rel(E,f`x,f`y)|] ==> \
\ f:mono(D,E)";
-by (fast_tac(ZF_cs addSIs prems) 1);
+by (fast_tac(!claset addSIs prems) 1);
val monoI = result();
val prems = goal Limit.thy
@@ -692,7 +684,7 @@
\ !!x y. [|rel(D,x,y); x:set(D); y:set(D)|] ==> rel(E,f`x,f`y); \
\ !!X. chain(D,X) ==> f`lub(D,X) = lub(E,lam n:nat. f`(X`n))|] ==> \
\ f:cont(D,E)";
-by (fast_tac(ZF_cs addSIs prems) 1);
+by (fast_tac(!claset addSIs prems) 1);
val contI = result();
val prems = goal Limit.thy
@@ -730,8 +722,8 @@
val prems = goalw Limit.thy [] (* mono_chain *)
"[|f:mono(D,E); chain(D,X)|] ==> chain(E,lam n:nat. f`(X`n))";
by (rewtac chain_def);
-by (simp_tac arith_ss 1);
-by (safe_tac lemmas_cs);
+by (Simp_tac 1);
+by (safe_tac (!claset));
by (rtac lam_type 1);
by (rtac mono_map 1);
by (resolve_tac prems 1);
@@ -760,13 +752,13 @@
val prems = goalw Limit.thy [set_def,cf_def]
"!!z. f:set(cf(D,E)) ==> f:cont(D,E)";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val in_cf = result();
val cf_cont = result();
val prems = goalw Limit.thy [set_def,cf_def] (* Non-trivial with relation *)
"!!z. f:cont(D,E) ==> f:set(cf(D,E))";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val cont_cf = result();
(* rel_cf originally an equality. Now stated as two rules. Seemed easiest.
@@ -776,14 +768,14 @@
"[|!!x. x:set(D) ==> rel(E,f`x,g`x); f:cont(D,E); g:cont(D,E)|] ==> \
\ rel(cf(D,E),f,g)";
by (rtac rel_I 1);
-by (simp_tac (ZF_ss addsimps [cf_def]) 1);
-by (safe_tac lemmas_cs);
+by (simp_tac (!simpset addsimps [cf_def]) 1);
+by (safe_tac (!claset));
brr prems 1;
val rel_cfI = result();
val prems = goalw Limit.thy [rel_def,cf_def]
"!!z. [|rel(cf(D,E),f,g); x:set(D)|] ==> rel(E,f`x,g`x)";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val rel_cf = result();
(*----------------------------------------------------------------------*)
@@ -797,7 +789,7 @@
by (rtac apply_type 1);
by (resolve_tac prems 2);
by (REPEAT(ares_tac([cont_fun,in_cf,chain_in]@prems) 1));
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (REPEAT(ares_tac([rel_cf,chain_rel]@prems) 1));
val chain_cf = result();
@@ -805,8 +797,8 @@
"[|chain(cf(D,E),X); chain(D,Xa); cpo(D); cpo(E) |] ==> \
\ matrix(E,lam x:nat. lam xa:nat. X`x`(Xa`xa))";
by (rtac matrix_chainI 1);
-by (asm_simp_tac ZF_ss 1);
-by (asm_simp_tac ZF_ss 2);
+by (Asm_simp_tac 1);
+by (Asm_simp_tac 2);
by (rtac chainI 1);
by (rtac lam_type 1);
by (rtac apply_type 1);
@@ -814,7 +806,7 @@
by (REPEAT(ares_tac prems 1));
by (rtac chain_in 1);
by (REPEAT(ares_tac prems 1));
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (rtac cont_mono 1);
by (rtac (chain_in RS cf_cont) 1);
brr prems 1;
@@ -826,7 +818,7 @@
by (REPEAT(ares_tac prems 1));
by (rtac chain_in 1);
by (REPEAT(ares_tac prems 1));
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (rtac rel_cf 1);
brr (chain_in::chain_rel::prems) 1;
by (rtac lam_type 1);
@@ -844,24 +836,24 @@
by (rtac contI 1);
by (rtac lam_type 1);
by (REPEAT(ares_tac((chain_cf RS cpo_lub RS islub_in)::prems) 1));
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (rtac dominate_islub 1);
by (REPEAT(ares_tac((chain_cf RS cpo_lub)::prems) 2));
by (rtac dominateI 1);
by (assume_tac 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (REPEAT(ares_tac ((chain_in RS cf_cont RS cont_mono)::prems) 1));
by (REPEAT(ares_tac ((chain_cf RS chain_fun)::prems) 1));
by (stac beta 1);
by (REPEAT(ares_tac((cpo_lub RS islub_in)::prems) 1));
-by (asm_simp_tac(ZF_ss addsimps[hd prems RS chain_in RS cf_cont RS cont_lub]) 1);
+by (asm_simp_tac(!simpset addsimps[hd prems RS chain_in RS cf_cont RS cont_lub]) 1);
by (forward_tac[hd prems RS matrix_lemma RS lub_matrix_diag]1);
brr prems 1;
-by (asm_full_simp_tac ZF_ss 1);
-by (asm_simp_tac(ZF_ss addsimps[chain_in RS beta]) 1);
+by (Asm_full_simp_tac 1);
+by (asm_simp_tac(!simpset addsimps[chain_in RS beta]) 1);
by (dtac (hd prems RS matrix_lemma RS lub_matrix_diag_sym) 1);
brr prems 1;
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val chain_cf_lub_cont = result();
val prems = goal Limit.thy (* islub_cf *)
@@ -872,21 +864,21 @@
by (rtac (chain_cf_lub_cont RS cont_cf) 1);
brr prems 1;
by (rtac rel_cfI 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (dtac (hd(tl(tl prems)) RSN(2,hd prems RS chain_cf RS cpo_lub RS islub_ub)) 1);
by (assume_tac 1);
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
brr(cf_cont::chain_in::prems) 1;
brr(cont_cf::chain_cf_lub_cont::prems) 1;
by (rtac rel_cfI 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (forward_tac[hd(tl(tl prems)) RSN(2,hd prems RS chain_cf RS cpo_lub RS
islub_least)]1);
by (assume_tac 2);
brr (chain_cf_lub_cont::isubD1::cf_cont::prems) 2;
by (rtac isubI 1);
brr((cf_cont RS cont_fun RS apply_type)::[isubD1]) 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (etac (isubD2 RS rel_cf) 1);
brr [] 1;
val islub_cf = result();
@@ -924,23 +916,19 @@
brr (cpo_lub::islub_cf::cpo_cf::prems) 1;
val lub_cf = result();
-val const_fun = prove_goal ZF.thy
- "y:set(E) ==> (lam x:set(D).y): set(D)->set(E)"
- (fn prems => [rtac lam_type 1,rtac (hd prems) 1]);
-
val prems = goal Limit.thy (* const_cont *)
"[|y:set(E); cpo(D); cpo(E)|] ==> (lam x:set(D).y) : cont(D,E)";
by (rtac contI 1);
-by (asm_simp_tac ZF_ss 2);
-brr(const_fun::cpo_refl::prems) 1;
-by (asm_simp_tac(ZF_ss addsimps(chain_in::(cpo_lub RS islub_in)::
+by (Asm_simp_tac 2);
+brr(lam_type::cpo_refl::prems) 1;
+by (asm_simp_tac(!simpset addsimps(chain_in::(cpo_lub RS islub_in)::
lub_const::prems)) 1);
val const_cont = result();
val prems = goal Limit.thy (* cf_least *)
"[|cpo(D); pcpo(E); y:cont(D,E)|]==>rel(cf(D,E),(lam x:set(D).bot(E)),y)";
by (rtac rel_cfI 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(bot_least::bot_in::apply_type::cont_fun::const_cont::
cpo_cf::(prems@[pcpo_cpo])) 1;
val cf_least = result();
@@ -964,14 +952,14 @@
(*----------------------------------------------------------------------*)
val id_thm = prove_goalw Perm.thy [id_def] "x:X ==> (id(X)`x) = x"
- (fn prems => [simp_tac(ZF_ss addsimps prems) 1]);
+ (fn prems => [simp_tac(!simpset addsimps prems) 1]);
val prems = goal Limit.thy (* id_cont *)
"cpo(D) ==> id(set(D)):cont(D,D)";
by (rtac contI 1);
by (rtac id_type 1);
-by (asm_simp_tac (ZF_ss addsimps[id_thm]) 1);
-by (asm_simp_tac(ZF_ss addsimps(id_thm::(cpo_lub RS islub_in)::
+by (asm_simp_tac (!simpset addsimps[id_thm]) 1);
+by (asm_simp_tac(!simpset addsimps(id_thm::(cpo_lub RS islub_in)::
chain_in::(chain_fun RS eta)::prems)) 1);
val id_cont = result();
@@ -988,7 +976,7 @@
by (stac comp_cont_apply 1);
by (stac cont_lub 4);
by (stac cont_lub 6);
-by (asm_full_simp_tac(ZF_ss addsimps (* RS: new subgoals contain unknowns *)
+by (asm_full_simp_tac(!simpset addsimps (* RS: new subgoals contain unknowns *)
[hd prems RS (hd(tl prems) RS comp_cont_apply),chain_in]) 8);
brr((cpo_lub RS islub_in)::cont_chain::prems) 1;
val comp_pres_cont = result();
@@ -1008,7 +996,7 @@
"[| chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(E)|] ==> \
\ chain(cf(D,E),lam n:nat. X`n O Y`n)";
by (rtac chainI 1);
-by (asm_simp_tac arith_ss 2);
+by (Asm_simp_tac 2);
by (rtac rel_cfI 2);
by (stac comp_cont_apply 2);
by (stac comp_cont_apply 5);
@@ -1027,51 +1015,35 @@
brr(comp_fun::(cf_cont RS cont_fun)::(cpo_lub RS islub_in)::cpo_cf::
chain_cf_comp::prems) 1;
by (cut_facts_tac[hd prems,hd(tl prems)]1);
-by (asm_simp_tac(ZF_ss addsimps((chain_in RS cf_cont RSN(3,chain_in RS
+by (asm_simp_tac(!simpset addsimps((chain_in RS cf_cont RSN(3,chain_in RS
cf_cont RS comp_cont_apply))::(tl(tl prems)))) 1);
by (stac comp_cont_apply 1);
brr((cpo_lub RS islub_in RS cf_cont)::cpo_cf::prems) 1;
-by (asm_simp_tac(ZF_ss addsimps(lub_cf::
+by (asm_simp_tac(!simpset addsimps(lub_cf::
(hd(tl prems)RS chain_cf RSN(2,hd prems RS chain_in RS cf_cont RS cont_lub))::
(hd(tl prems) RS chain_cf RS cpo_lub RS islub_in)::prems)) 1);
by (cut_inst_tac[("M","lam xa:nat. lam xb:nat. X`xa`(Y`xb`x)")]
lub_matrix_diag 1);
-by (asm_full_simp_tac ZF_ss 3);
+by (Asm_full_simp_tac 3);
by (rtac matrix_chainI 1);
-by (asm_simp_tac ZF_ss 1);
-by (asm_simp_tac ZF_ss 2);
+by (Asm_simp_tac 1);
+by (Asm_simp_tac 2);
by (forward_tac[hd(tl prems) RSN(2,(hd prems RS chain_in RS cf_cont) RS
(chain_cf RSN(2,cont_chain)))]1); (* Here, Isabelle was a bitch! *)
-by (asm_full_simp_tac ZF_ss 2);
+by (Asm_full_simp_tac 2);
by (assume_tac 1);
by (rtac chain_cf 1);
brr((cont_fun RS apply_type)::(chain_in RS cf_cont)::lam_type::prems) 1;
val comp_lubs = result();
(*----------------------------------------------------------------------*)
-(* A couple (out of many) theorems about arithmetic. *)
-(*----------------------------------------------------------------------*)
-
-val prems = goal Arith.thy (* le_cases *)
- "[|m:nat; n:nat|] ==> m le n | n le m";
-by (safe_tac lemmas_cs);
-brr((not_le_iff_lt RS iffD1 RS leI)::nat_into_Ord::prems) 1;
-val le_cases = result();
-
-val prems = goal Arith.thy (* lt_cases *)
- "[|m:nat; n:nat|] ==> m < n | n le m";
-by (safe_tac lemmas_cs);
-brr((not_le_iff_lt RS iffD1)::nat_into_Ord::prems) 1;
-val lt_cases = result();
-
-(*----------------------------------------------------------------------*)
(* Theorems about projpair. *)
(*----------------------------------------------------------------------*)
val prems = goalw Limit.thy [projpair_def] (* projpairI *)
"!!x. [| e:cont(D,E); p:cont(E,D); p O e = id(set(D)); \
\ rel(cf(E,E))(e O p)(id(set(E)))|] ==> projpair(D,E,e,p)";
-by (fast_tac FOL_cs 1);
+by (Fast_tac 1);
val projpairI = result();
val prems = goalw Limit.thy [projpair_def] (* projpairE *)
@@ -1079,7 +1051,7 @@
\ [| e:cont(D,E); p:cont(E,D); p O e = id(set(D)); \
\ rel(cf(E,E))(e O p)(id(set(E)))|] ==> Q |] ==> Q";
by (rtac (hd(tl prems)) 1);
-by (REPEAT(asm_simp_tac(FOL_ss addsimps[hd prems]) 1));
+by (REPEAT(asm_simp_tac(!simpset addsimps[hd prems]) 1));
val projpairE = result();
val prems = goal Limit.thy (* projpair_e_cont *)
@@ -1174,22 +1146,22 @@
(* First some existentials are instantiated. *)
by (resolve_tac prems 4);
by (resolve_tac prems 4);
-by (asm_simp_tac FOL_ss 4);
+by (Asm_simp_tac 4);
brr(cpo_cf::cpo_refl::cont_cf::projpair_e_cont::prems) 1;
by (rtac lemma1 1);
brr prems 1;
-by (asm_simp_tac FOL_ss 1);
+by (Asm_simp_tac 1);
brr(cpo_cf::cpo_refl::cont_cf::(contl @ prems)) 1;
by (rtac cpo_antisym 1);
by (rtac lemma2 2);
(* First some existentials are instantiated. *)
by (resolve_tac prems 4);
by (resolve_tac prems 4);
-by (asm_simp_tac FOL_ss 4);
+by (Asm_simp_tac 4);
brr(cpo_cf::cpo_refl::cont_cf::projpair_p_cont::prems) 1;
by (rtac lemma2 1);
brr prems 1;
-by (asm_simp_tac FOL_ss 1);
+by (Asm_simp_tac 1);
brr(cpo_cf::cpo_refl::cont_cf::(contl @ prems)) 1;
val projpair_unique = result();
@@ -1209,7 +1181,7 @@
val embI = prove_goalw Limit.thy [emb_def]
"!!x. projpair(D,E,e,p) ==> emb(D,E,e)"
- (fn prems => [fast_tac FOL_cs 1]);
+ (fn prems => [Fast_tac 1]);
val prems = goal Limit.thy (* Rp_unique *)
"[|projpair(D,E,e,p); cpo(D); cpo(E)|] ==> Rp(D,E,e) = p";
@@ -1230,7 +1202,7 @@
val id_apply = prove_goalw Perm.thy [id_def]
"!!z. x:A ==> id(A)`x = x"
- (fn prems => [asm_simp_tac ZF_ss 1]);
+ (fn prems => [Asm_simp_tac 1]);
val prems = goal Limit.thy (* embRp_eq_thm *)
"[|emb(D,E,e); x:set(D); cpo(D); cpo(E)|] ==> Rp(D,E,e)`(e`x) = x";
@@ -1247,7 +1219,7 @@
val prems = goalw Limit.thy [projpair_def] (* projpair_id *)
"cpo(D) ==> projpair(D,D,id(set(D)),id(set(D)))";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
brr(id_cont::id_comp::id_type::prems) 1;
by (stac id_comp 1); (* Matches almost anything *)
brr(id_cont::id_type::cpo_refl::cpo_cf::cont_cf::prems) 1;
@@ -1274,7 +1246,7 @@
val prems = goalw Limit.thy [projpair_def] (* lemma *)
"[|emb(D,D',e); emb(D',E,e'); cpo(D); cpo(D'); cpo(E)|] ==> \
\ projpair(D,E,e' O e,(Rp(D,D',e)) O (Rp(D',E,e')))";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
brr(comp_pres_cont::Rp_cont::emb_cont::prems) 1;
by (rtac (comp_assoc RS subst) 1);
by (res_inst_tac[("t1","e'")](comp_assoc RS ssubst) 1);
@@ -1310,25 +1282,25 @@
val prems = goalw Limit.thy [set_def,iprod_def] (* iprodI *)
"!!z. x:(PROD n:nat. set(DD`n)) ==> x:set(iprod(DD))";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val iprodI = result();
(* Proof with non-reflexive relation approach:
by (rtac CollectI 1);
by (rtac domainI 1);
by (rtac CollectI 1);
-by (simp_tac(ZF_ss addsimps prems) 1);
+by (simp_tac(!simpset addsimps prems) 1);
by (rtac (hd prems) 1);
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
by (rtac ballI 1);
by (dtac ((hd prems) RS apply_type) 1);
by (etac CollectE 1);
by (assume_tac 1);
by (rtac rel_I 1);
by (rtac CollectI 1);
-by (fast_tac(ZF_cs addSIs prems) 1);
+by (fast_tac(!claset addSIs prems) 1);
by (rtac ballI 1);
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
by (dtac ((hd prems) RS apply_type) 1);
by (etac CollectE 1);
by (assume_tac 1);
@@ -1336,7 +1308,7 @@
val prems = goalw Limit.thy [set_def,iprod_def] (* iprodE *)
"!!z. x:set(iprod(DD)) ==> x:(PROD n:nat. set(DD`n))";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val iprodE = result();
(* Contains typing conditions in contrast to HOL-ST *)
@@ -1345,16 +1317,16 @@
"[|!!n. n:nat ==> rel(DD`n,f`n,g`n); f:(PROD n:nat. set(DD`n)); \
\ g:(PROD n:nat. set(DD`n))|] ==> rel(iprod(DD),f,g)";
by (rtac rel_I 1);
-by (simp_tac ZF_ss 1);
-by (safe_tac lemmas_cs);
+by (Simp_tac 1);
+by (safe_tac (!claset));
brr prems 1;
val rel_iprodI = result();
val prems = goalw Limit.thy [iprod_def] (* rel_iprodE *)
"[|rel(iprod(DD),f,g); n:nat|] ==> rel(DD`n,f`n,g`n)";
by (cut_facts_tac[hd prems RS rel_E]1);
-by (asm_full_simp_tac ZF_ss 1);
-by (safe_tac lemmas_cs);
+by (Asm_full_simp_tac 1);
+by (safe_tac (!claset));
by (etac bspec 1);
by (resolve_tac prems 1);
val rel_iprodE = result();
@@ -1363,42 +1335,42 @@
probably not needed in Isabelle, wait and see. *)
val prems = goalw Limit.thy [chain_def] (* chain_iprod *)
- "[|chain(iprod(DD),X);!!n. n:nat ==> cpo(DD`n); n:nat|] ==> \
+ "[|chain(iprod(DD),X); !!n. n:nat ==> cpo(DD`n); n:nat|] ==> \
\ chain(DD`n,lam m:nat.X`m`n)";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac lam_type 1);
by (rtac apply_type 1);
by (rtac iprodE 1);
by (etac (hd prems RS conjunct1 RS apply_type) 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(arith_ss addsimps prems) 1);
+by (asm_simp_tac(!simpset addsimps prems) 1);
by (rtac rel_iprodE 1);
-by (asm_simp_tac (arith_ss addsimps prems) 1);
+by (asm_simp_tac (!simpset addsimps prems) 1);
by (resolve_tac prems 1);
val chain_iprod = result();
val prems = goalw Limit.thy [islub_def,isub_def] (* islub_iprod *)
- "[|chain(iprod(DD),X);!!n. n:nat ==> cpo(DD`n)|] ==> \
+ "[|chain(iprod(DD),X); !!n. n:nat ==> cpo(DD`n)|] ==> \
\ islub(iprod(DD),X,lam n:nat. lub(DD`n,lam m:nat.X`m`n))";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac iprodI 1);
by (rtac lam_type 1);
brr((chain_iprod RS cpo_lub RS islub_in)::prems) 1;
by (rtac rel_iprodI 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
(* Here, HOL resolution is handy, Isabelle resolution bad. *)
by (res_inst_tac[("P","%t. rel(DD`na,t,lub(DD`na,lam x:nat. X`x`na))"),
("b1","%n. X`n`na")](beta RS subst) 1);
brr((chain_iprod RS cpo_lub RS islub_ub)::iprodE::chain_in::prems) 1;
brr(iprodI::lam_type::(chain_iprod RS cpo_lub RS islub_in)::prems) 1;
by (rtac rel_iprodI 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(islub_least::(chain_iprod RS cpo_lub)::prems) 1;
by (rewtac isub_def);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (etac (iprodE RS apply_type) 1);
by (assume_tac 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (dtac bspec 1);
by (etac rel_iprodE 2);
brr(lam_type::(chain_iprod RS cpo_lub RS islub_in)::iprodE::prems) 1;
@@ -1419,7 +1391,7 @@
val cpo_iprod = result();
val prems = goalw Limit.thy [islub_def,isub_def] (* lub_iprod *)
- "[|chain(iprod(DD),X);!!n. n:nat ==> cpo(DD`n)|] ==> \
+ "[|chain(iprod(DD),X); !!n. n:nat ==> cpo(DD`n)|] ==> \
\ lub(iprod(DD),X) = (lam n:nat. lub(DD`n,lam m:nat.X`m`n))";
brr((cpo_lub RS islub_unique)::islub_iprod::cpo_iprod::prems) 1;
val lub_iprod = result();
@@ -1432,15 +1404,15 @@
"[|set(D)<=set(E); \
\ !!x y. [|x:set(D); y:set(D)|] ==> rel(D,x,y)<->rel(E,x,y); \
\ !!X. chain(D,X) ==> lub(E,X) : set(D)|] ==> subcpo(D,E)";
-by (safe_tac lemmas_cs);
-by (asm_full_simp_tac(ZF_ss addsimps prems) 2);
-by (asm_simp_tac(ZF_ss addsimps prems) 2);
+by (safe_tac (!claset));
+by (asm_full_simp_tac(!simpset addsimps prems) 2);
+by (asm_simp_tac(!simpset addsimps prems) 2);
brr(prems@[subsetD]) 1;
val subcpoI = result();
val subcpo_subset = prove_goalw Limit.thy [subcpo_def]
"!!x. subcpo(D,E) ==> set(D)<=set(E)"
- (fn prems => [fast_tac FOL_cs 1]);
+ (fn prems => [Fast_tac 1]);
val subcpo_rel_eq = prove_goalw Limit.thy [subcpo_def]
" [|subcpo(D,E); x:set(D); y:set(D)|] ==> rel(D,x,y)<->rel(E,x,y)"
@@ -1488,16 +1460,16 @@
"[|subcpo(D,E); cpo(E)|] ==> cpo(D)";
brr[cpoI,poI]1;
(* Changing the order of the assumptions, otherwise full_simp doesn't work. *)
-by (asm_full_simp_tac(ZF_ss addsimps[hd prems RS subcpo_rel_eq]) 1);
+by (asm_full_simp_tac(!simpset addsimps[hd prems RS subcpo_rel_eq]) 1);
brr(cpo_refl::(hd prems RS subcpo_subset RS subsetD)::prems) 1;
by (dtac (imp_refl RS mp) 1);
by (dtac (imp_refl RS mp) 1);
-by (asm_full_simp_tac(ZF_ss addsimps[hd prems RS subcpo_rel_eq]) 1);
+by (asm_full_simp_tac(!simpset addsimps[hd prems RS subcpo_rel_eq]) 1);
brr(cpo_trans::(hd prems RS subcpo_subset RS subsetD)::prems) 1;
(* Changing the order of the assumptions, otherwise full_simp doesn't work. *)
by (dtac (imp_refl RS mp) 1);
by (dtac (imp_refl RS mp) 1);
-by (asm_full_simp_tac(ZF_ss addsimps[hd prems RS subcpo_rel_eq]) 1);
+by (asm_full_simp_tac(!simpset addsimps[hd prems RS subcpo_rel_eq]) 1);
brr(cpo_antisym::(hd prems RS subcpo_subset RS subsetD)::prems) 1;
brr(islub_subcpo::prems) 1;
val subcpo_cpo = result();
@@ -1513,7 +1485,7 @@
val prems = goalw Limit.thy [set_def,mkcpo_def] (* mkcpoI *)
"!!z. [|x:set(D); P(x)|] ==> x:set(mkcpo(D,P))";
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
brr(conjI::prems) 1;
val mkcpoI = result();
@@ -1523,41 +1495,41 @@
by (rtac domainI 1);
by (rtac CollectI 1);
(* Now, work on subgoal 2 (and 3) to instantiate unknown. *)
-by (simp_tac ZF_ss 2);
+by (Simp_tac 2);
by (rtac conjI 2);
by (rtac conjI 3);
by (resolve_tac prems 3);
-by (simp_tac(ZF_ss addsimps [rewrite_rule[set_def](hd prems)]) 1);
+by (simp_tac(!simpset addsimps [rewrite_rule[set_def](hd prems)]) 1);
by (resolve_tac prems 1);
by (rtac cpo_refl 1);
by (resolve_tac prems 1);
by (rtac rel_I 1);
by (rtac CollectI 1);
-by (fast_tac(ZF_cs addSIs [rewrite_rule[set_def](hd prems)]) 1);
-by (simp_tac ZF_ss 1);
+by (fast_tac(!claset addSIs [rewrite_rule[set_def](hd prems)]) 1);
+by (Simp_tac 1);
brr(conjI::cpo_refl::prems) 1;
*)
val prems = goalw Limit.thy [set_def,mkcpo_def] (* mkcpoD1 *)
"!!z. x:set(mkcpo(D,P))==> x:set(D)";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val mkcpoD1 = result();
val prems = goalw Limit.thy [set_def,mkcpo_def] (* mkcpoD2 *)
"!!z. x:set(mkcpo(D,P))==> P(x)";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val mkcpoD2 = result();
val prems = goalw Limit.thy [rel_def,mkcpo_def] (* rel_mkcpoE *)
"!!a. rel(mkcpo(D,P),x,y) ==> rel(D,x,y)";
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val rel_mkcpoE = result();
val rel_mkcpo = prove_goalw Limit.thy [mkcpo_def,rel_def,set_def]
"!!z. [|x:set(D); y:set(D)|] ==> rel(mkcpo(D,P),x,y) <-> rel(D,x,y)"
- (fn prems => [asm_simp_tac ZF_ss 1]);
+ (fn prems => [Asm_simp_tac 1]);
-(* The HOL proof is simpler, problems due to cpos as purely in ZF. *)
+(* The HOL proof is simpler, problems due to cpos as purely in upair. *)
(* And chains as set functions. *)
val prems = goal Limit.thy (* chain_mkcpo *)
@@ -1593,17 +1565,17 @@
val prems = goalw Limit.thy [emb_chain_def] (* emb_chainI *)
"[|!!n. n:nat ==> cpo(DD`n); \
\ !!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n)|] ==> emb_chain(DD,ee)";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
brr prems 1;
val emb_chainI = result();
val emb_chain_cpo = prove_goalw Limit.thy [emb_chain_def]
"!!x. [|emb_chain(DD,ee); n:nat|] ==> cpo(DD`n)"
- (fn prems => [fast_tac ZF_cs 1]);
+ (fn prems => [Fast_tac 1]);
val emb_chain_emb = prove_goalw Limit.thy [emb_chain_def]
"!!x. [|emb_chain(DD,ee); n:nat|] ==> emb(DD`n,DD`succ(n),ee`n)"
- (fn prems => [fast_tac ZF_cs 1]);
+ (fn prems => [Fast_tac 1]);
(*----------------------------------------------------------------------*)
(* Dinf, the inverse Limit. *)
@@ -1627,7 +1599,7 @@
val prems = goalw Limit.thy [Dinf_def] (* DinfD2 *)
"[|x:set(Dinf(DD,ee)); n:nat|] ==> \
\ Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n";
-by (asm_simp_tac(ZF_ss addsimps[(hd prems RS mkcpoD2),hd(tl prems)]) 1);
+by (asm_simp_tac(!simpset addsimps[(hd prems RS mkcpoD2),hd(tl prems)]) 1);
val DinfD2 = result();
val Dinf_eq = DinfD2;
@@ -1666,11 +1638,11 @@
by (rtac ballI 1);
by (stac lub_iprod 1);
brr(chain_Dinf::(hd prems RS emb_chain_cpo)::[]) 1;
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (stac (Rp_cont RS cont_lub) 1);
brr(emb_chain_cpo::emb_chain_emb::nat_succI::chain_iprod::chain_Dinf::prems) 1;
(* Useful simplification, ugly in HOL. *)
-by (asm_simp_tac(arith_ss addsimps(DinfD2::chain_in::[])) 1);
+by (asm_simp_tac(!simpset addsimps(DinfD2::chain_in::[])) 1);
brr(cpo_iprod::emb_chain_cpo::prems) 1;
val subcpo_Dinf = result();
@@ -1699,22 +1671,23 @@
val prems = goalw Limit.thy [e_less_def] (* e_less_eq *)
"!!x. m:nat ==> e_less(DD,ee,m,m) = id(set(DD`m))";
-by (asm_simp_tac (arith_ss addsimps[diff_self_eq_0]) 1);
+by (asm_simp_tac (!simpset addsimps[diff_self_eq_0]) 1);
val e_less_eq = result();
(* ARITH_CONV proves the following in HOL. Would like something similar
in Isabelle/ZF. *)
-goalw Arith.thy [] (* lemma_succ_sub *)
+goal Arith.thy (* lemma_succ_sub *)
"!!z. [|n:nat; m:nat|] ==> succ(m#+n)#-m = succ(n)";
(*Uses add_succ_right the wrong way round!*)
-by (asm_simp_tac(nat_ss addsimps [add_succ_right RS sym, diff_add_inverse]) 1);
+by (asm_simp_tac
+ (simpset_of"Nat" addsimps [add_succ_right RS sym, diff_add_inverse]) 1);
val lemma_succ_sub = result();
val prems = goalw Limit.thy [e_less_def] (* e_less_add *)
"!!x. [|m:nat; k:nat|] ==> \
\ e_less(DD,ee,m,succ(m#+k)) = (ee`(m#+k))O(e_less(DD,ee,m,m#+k))";
-by (asm_simp_tac (arith_ss addsimps [lemma_succ_sub,diff_add_inverse]) 1);
+by (asm_simp_tac (!simpset addsimps [lemma_succ_sub,diff_add_inverse]) 1);
val e_less_add = result();
(* Again, would like more theorems about arithmetic. *)
@@ -1723,16 +1696,16 @@
val add1 = prove_goal Limit.thy
"!!x. n:nat ==> succ(n) = n #+ 1"
(fn prems =>
- [asm_simp_tac (arith_ss addsimps[add_succ_right,add_0_right]) 1]);
+ [asm_simp_tac (!simpset addsimps[add_succ_right,add_0_right]) 1]);
val prems = goal Limit.thy (* succ_sub1 *)
"x:nat ==> 0 < x --> succ(x#-1)=x";
by (res_inst_tac[("n","x")]nat_induct 1);
by (resolve_tac prems 1);
-by (fast_tac lt_cs 1);
-by (safe_tac lemmas_cs);
-by (asm_simp_tac arith_ss 1);
-by (asm_simp_tac arith_ss 1);
+by (Fast_tac 1);
+by (safe_tac (!claset));
+by (Asm_simp_tac 1);
+by (Asm_simp_tac 1);
val succ_sub1 = result();
val prems = goal Limit.thy (* succ_le_pos *)
@@ -1740,23 +1713,24 @@
by (res_inst_tac[("n","m")]nat_induct 1);
by (resolve_tac prems 1);
by (rtac impI 1);
-by (asm_full_simp_tac(arith_ss addsimps prems) 1);
-by (safe_tac lemmas_cs);
-by (asm_full_simp_tac(arith_ss addsimps prems) 1); (* Surprise, surprise. *)
+by (asm_full_simp_tac(!simpset addsimps prems) 1);
+by (safe_tac (!claset));
+by (asm_full_simp_tac(!simpset addsimps prems) 1); (* Surprise, surprise. *)
val succ_le_pos = result();
-val prems = goal Limit.thy (* lemma_le_exists *)
+goal Limit.thy (* lemma_le_exists *)
"!!z. [|n:nat; m:nat|] ==> m le n --> (EX k:nat. n = m #+ k)";
by (res_inst_tac[("n","m")]nat_induct 1);
by (assume_tac 1);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (rtac bexI 1);
by (rtac (add_0 RS sym) 1);
by (assume_tac 1);
-by (asm_full_simp_tac arith_ss 1);
-(* Great, by luck I found lt_cs. Such cs's and ss's should be documented. *)
-by (fast_tac lt_cs 1);
-by (asm_simp_tac (nat_ss addsimps[add_succ, add_succ_right RS sym]) 1);
+by (Asm_full_simp_tac 1);
+(* Great, by luck I found le_cs. Such cs's and ss's should be documented. *)
+by (fast_tac le_cs 1);
+by (asm_simp_tac
+ (simpset_of"Nat" addsimps[add_succ, add_succ_right RS sym]) 1);
by (rtac bexI 1);
by (stac (succ_sub1 RS mp) 1);
(* Instantiation. *)
@@ -1765,11 +1739,11 @@
by (rtac (succ_le_pos RS mp) 1);
by (assume_tac 3); (* Instantiation *)
brr[]1;
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
val lemma_le_exists = result();
val prems = goal Limit.thy (* le_exists *)
- "[|m le n;!!x. [|n=m#+x; x:nat|] ==> Q; m:nat; n:nat|] ==> Q";
+ "[|m le n; !!x. [|n=m#+x; x:nat|] ==> Q; m:nat; n:nat|] ==> Q";
by (rtac (lemma_le_exists RS mp RS bexE) 1);
by (rtac (hd(tl prems)) 4);
by (assume_tac 4);
@@ -1781,7 +1755,7 @@
\ e_less(DD,ee,m,succ(n)) = ee`n O e_less(DD,ee,m,n)";
by (rtac le_exists 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(ZF_ss addsimps(e_less_add::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_less_add::prems)) 1);
brr prems 1;
val e_less_le = result();
@@ -1789,13 +1763,13 @@
val prems = goal Limit.thy (* e_less_succ *)
"m:nat ==> e_less(DD,ee,m,succ(m)) = ee`m O id(set(DD`m))";
-by (asm_simp_tac(arith_ss addsimps(e_less_le::e_less_eq::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_less_le::e_less_eq::prems)) 1);
val e_less_succ = result();
val prems = goal Limit.thy (* e_less_succ_emb *)
"[|!!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n); m:nat|] ==> \
\ e_less(DD,ee,m,succ(m)) = ee`m";
-by (asm_simp_tac(arith_ss addsimps(e_less_succ::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_less_succ::prems)) 1);
by (stac comp_id 1);
brr(emb_cont::cont_fun::refl::prems) 1;
val e_less_succ_emb = result();
@@ -1808,9 +1782,9 @@
\ emb(DD`m,DD`(m#+k),e_less(DD,ee,m,m#+k))";
by (res_inst_tac[("n","k")]nat_induct 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(ZF_ss addsimps(add_0_right::e_less_eq::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(add_0_right::e_less_eq::prems)) 1);
brr(emb_id::emb_chain_cpo::prems) 1;
-by (asm_simp_tac(ZF_ss addsimps(add_succ_right::e_less_add::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(add_succ_right::e_less_add::prems)) 1);
brr(emb_comp::emb_chain_emb::emb_chain_cpo::add_type::nat_succI::prems) 1;
val emb_e_less_add = result();
@@ -1820,13 +1794,13 @@
(* same proof as e_less_le *)
by (rtac le_exists 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(ZF_ss addsimps(emb_e_less_add::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(emb_e_less_add::prems)) 1);
brr prems 1;
val emb_e_less = result();
val comp_mono_eq = prove_goal Limit.thy
"!!z.[|f=f'; g=g'|] ==> f O g = f' O g'"
- (fn prems => [asm_simp_tac ZF_ss 1]);
+ (fn prems => [Asm_simp_tac 1]);
(* Typing, typing, typing, three irritating assumptions. Extra theorems
needed in proof, but no real difficulty. *)
@@ -1868,14 +1842,14 @@
val prems = goalw Limit.thy [e_gr_def] (* e_gr_eq *)
"!!x. m:nat ==> e_gr(DD,ee,m,m) = id(set(DD`m))";
-by (asm_simp_tac (arith_ss addsimps[diff_self_eq_0]) 1);
+by (asm_simp_tac (!simpset addsimps[diff_self_eq_0]) 1);
val e_gr_eq = result();
val prems = goalw Limit.thy [e_gr_def] (* e_gr_add *)
"!!x. [|n:nat; k:nat|] ==> \
\ e_gr(DD,ee,succ(n#+k),n) = \
\ e_gr(DD,ee,n#+k,n) O Rp(DD`(n#+k),DD`succ(n#+k),ee`(n#+k))";
-by (asm_simp_tac (arith_ss addsimps [lemma_succ_sub,diff_add_inverse]) 1);
+by (asm_simp_tac (!simpset addsimps [lemma_succ_sub,diff_add_inverse]) 1);
val e_gr_add = result();
val prems = goal Limit.thy (* e_gr_le *)
@@ -1883,14 +1857,14 @@
\ e_gr(DD,ee,succ(m),n) = e_gr(DD,ee,m,n) O Rp(DD`m,DD`succ(m),ee`m)";
by (rtac le_exists 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(ZF_ss addsimps(e_gr_add::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_gr_add::prems)) 1);
brr prems 1;
val e_gr_le = result();
val prems = goal Limit.thy (* e_gr_succ *)
"m:nat ==> \
\ e_gr(DD,ee,succ(m),m) = id(set(DD`m)) O Rp(DD`m,DD`succ(m),ee`m)";
-by (asm_simp_tac(arith_ss addsimps(e_gr_le::e_gr_eq::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_gr_le::e_gr_eq::prems)) 1);
val e_gr_succ = result();
(* Cpo asm's due to THE uniqueness. *)
@@ -1898,7 +1872,7 @@
val prems = goal Limit.thy (* e_gr_succ_emb *)
"[|emb_chain(DD,ee); m:nat|] ==> \
\ e_gr(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)";
-by (asm_simp_tac(arith_ss addsimps(e_gr_succ::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_gr_succ::prems)) 1);
by (stac id_comp 1);
brr(Rp_cont::cont_fun::refl::emb_chain_cpo::emb_chain_emb::nat_succI::prems) 1;
val e_gr_succ_emb = result();
@@ -1908,8 +1882,8 @@
\ e_gr(DD,ee,n#+k,n): set(DD`(n#+k))->set(DD`n)";
by (res_inst_tac[("n","k")]nat_induct 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(ZF_ss addsimps(add_0_right::e_gr_eq::id_type::prems)) 1);
-by (asm_simp_tac(ZF_ss addsimps(add_succ_right::e_gr_add::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(add_0_right::e_gr_eq::id_type::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(add_succ_right::e_gr_add::prems)) 1);
brr(comp_fun::Rp_cont::cont_fun::emb_chain_emb::emb_chain_cpo::add_type::
nat_succI::prems) 1;
val e_gr_fun_add = result();
@@ -1919,7 +1893,7 @@
\ e_gr(DD,ee,m,n): set(DD`m)->set(DD`n)";
by (rtac le_exists 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(ZF_ss addsimps(e_gr_fun_add::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_gr_fun_add::prems)) 1);
brr prems 1;
val e_gr_fun = result();
@@ -1965,16 +1939,16 @@
\ n le m --> e_gr(DD,ee,m,n):cont(DD`m,DD`n)";
by (res_inst_tac[("n","m")]nat_induct 1);
by (resolve_tac prems 1);
-by (asm_full_simp_tac(ZF_ss addsimps
+by (asm_full_simp_tac(!simpset addsimps
(le0_iff::e_gr_eq::nat_0I::prems)) 1);
brr(impI::id_cont::emb_chain_cpo::nat_0I::prems) 1;
-by (asm_full_simp_tac(ZF_ss addsimps[le_succ_iff]) 1);
+by (asm_full_simp_tac(!simpset addsimps[le_succ_iff]) 1);
by (etac disjE 1);
by (etac impE 1);
by (assume_tac 1);
-by (asm_simp_tac(ZF_ss addsimps(e_gr_le::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_gr_le::prems)) 1);
brr(comp_pres_cont::Rp_cont::emb_chain_cpo::emb_chain_emb::nat_succI::prems) 1;
-by (asm_simp_tac(ZF_ss addsimps(e_gr_eq::nat_succI::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_gr_eq::nat_succI::prems)) 1);
brr(id_cont::emb_chain_cpo::nat_succI::prems) 1;
val e_gr_cont_lemma = result();
@@ -1995,8 +1969,7 @@
by (res_inst_tac[("n","k")]nat_induct 1);
by (resolve_tac prems 1);
by (asm_full_simp_tac(ZF_ss addsimps
- (le0_iff::add_0_right::e_gr_eq::e_less_eq::(id_type RS id_comp)::prems)) 1);
-by (simp_tac(ZF_ss addsimps[le_succ_iff]) 1);
+ (le0_iff::add_0_right::e_gr_eq::e_less_eq::(id_type RS id_comp)::prems)) 1);by (simp_tac(ZF_ss addsimps[le_succ_iff]) 1);
by (rtac impI 1);
by (etac disjE 1);
by (etac impE 1);
@@ -2011,7 +1984,7 @@
brr((e_less_cont RS cont_fun)::add_type::add_le_self::refl::prems) 1;
by (asm_full_simp_tac(ZF_ss addsimps(e_gr_eq::nat_succI::add_type::prems)) 1);
by (stac id_comp 1);
-brr((e_less_cont RS cont_fun)::add_type::nat_succI::add_le_self::refl::prems) 1;
+brr((e_less_cont RS cont_fun)::add_type::nat_succI::add_le_self::refl::prems)1;
val e_less_e_gr_split_add = result();
(* Again considerably shorter, and easy to obtain from the previous thm. *)
@@ -2024,7 +1997,7 @@
by (resolve_tac prems 2);
by (res_inst_tac[("n","k")]nat_induct 1);
by (resolve_tac prems 1);
-by (asm_full_simp_tac(arith_ss addsimps
+by (asm_full_simp_tac(!simpset addsimps
(add_0_right::e_gr_eq::e_less_eq::(id_type RS id_comp)::prems)) 1);
by (simp_tac(ZF_ss addsimps[le_succ_iff]) 1);
by (rtac impI 1);
@@ -2040,15 +2013,16 @@
by (stac id_comp 1);
brr((e_less_cont RS cont_fun)::add_type::add_le_mono::nat_le_refl::refl::
prems) 1;
-by (asm_full_simp_tac(ZF_ss addsimps(e_less_eq::nat_succI::add_type::prems)) 1);
+by(asm_full_simp_tac(ZF_ss addsimps(e_less_eq::nat_succI::add_type::prems)) 1);
by (stac comp_id 1);
brr((e_gr_cont RS cont_fun)::add_type::nat_succI::add_le_self::refl::prems) 1;
val e_gr_e_less_split_add = result();
+
val prems = goalw Limit.thy [eps_def] (* emb_eps *)
"[|m le n; emb_chain(DD,ee); m:nat; n:nat|] ==> \
\ emb(DD`m,DD`n,eps(DD,ee,m,n))";
-by (asm_simp_tac(ZF_ss addsimps prems) 1);
+by (asm_simp_tac(!simpset addsimps prems) 1);
brr(emb_e_less::prems) 1;
val emb_eps = result();
@@ -2056,66 +2030,66 @@
"[|emb_chain(DD,ee); m:nat; n:nat|] ==> \
\ eps(DD,ee,m,n): set(DD`m)->set(DD`n)";
by (rtac (expand_if RS iffD2) 1);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
brr((e_less_cont RS cont_fun)::prems) 1;
brr((not_le_iff_lt RS iffD1 RS leI)::e_gr_fun::nat_into_Ord::prems) 1;
val eps_fun = result();
val eps_id = prove_goalw Limit.thy [eps_def]
"n:nat ==> eps(DD,ee,n,n) = id(set(DD`n))"
- (fn prems => [simp_tac(ZF_ss addsimps(e_less_eq::nat_le_refl::prems)) 1]);
+ (fn prems => [simp_tac(!simpset addsimps(e_less_eq::nat_le_refl::prems)) 1]);
val eps_e_less_add = prove_goalw Limit.thy [eps_def]
"[|m:nat; n:nat|] ==> eps(DD,ee,m,m#+n) = e_less(DD,ee,m,m#+n)"
- (fn prems => [simp_tac(ZF_ss addsimps(add_le_self::prems)) 1]);
+ (fn prems => [simp_tac(!simpset addsimps(add_le_self::prems)) 1]);
val eps_e_less = prove_goalw Limit.thy [eps_def]
"[|m le n; m:nat; n:nat|] ==> eps(DD,ee,m,n) = e_less(DD,ee,m,n)"
- (fn prems => [simp_tac(ZF_ss addsimps prems) 1]);
+ (fn prems => [simp_tac(!simpset addsimps prems) 1]);
val shift_asm = imp_refl RS mp;
val prems = goalw Limit.thy [eps_def] (* eps_e_gr_add *)
"[|n:nat; k:nat|] ==> eps(DD,ee,n#+k,n) = e_gr(DD,ee,n#+k,n)";
by (rtac (expand_if RS iffD2) 1);
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (etac leE 1);
(* Must control rewriting by instantiating a variable. *)
-by (asm_full_simp_tac(arith_ss addsimps
+by (asm_full_simp_tac(!simpset addsimps
((hd prems RS nat_into_Ord RS not_le_iff_lt RS iff_sym)::nat_into_Ord::
add_le_self::prems)) 1);
-by (asm_simp_tac(ZF_ss addsimps(e_less_eq::e_gr_eq::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_less_eq::e_gr_eq::prems)) 1);
val eps_e_gr_add = result();
val prems = goalw Limit.thy [] (* eps_e_gr *)
"[|n le m; m:nat; n:nat|] ==> eps(DD,ee,m,n) = e_gr(DD,ee,m,n)";
by (rtac le_exists 1);
by (resolve_tac prems 1);
-by (asm_simp_tac(ZF_ss addsimps(eps_e_gr_add::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(eps_e_gr_add::prems)) 1);
brr prems 1;
val eps_e_gr = result();
val prems = goal Limit.thy (* eps_succ_ee *)
"[|!!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n); m:nat|] ==> \
\ eps(DD,ee,m,succ(m)) = ee`m";
-by (asm_simp_tac(arith_ss addsimps(eps_e_less::le_succ_iff::e_less_succ_emb::
+by (asm_simp_tac(!simpset addsimps(eps_e_less::le_succ_iff::e_less_succ_emb::
prems)) 1);
val eps_succ_ee = result();
val prems = goal Limit.thy (* eps_succ_Rp *)
"[|emb_chain(DD,ee); m:nat|] ==> \
\ eps(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)";
-by (asm_simp_tac(arith_ss addsimps(eps_e_gr::le_succ_iff::e_gr_succ_emb::
+by (asm_simp_tac(!simpset addsimps(eps_e_gr::le_succ_iff::e_gr_succ_emb::
prems)) 1);
val eps_succ_Rp = result();
val prems = goal Limit.thy (* eps_cont *)
"[|emb_chain(DD,ee); m:nat; n:nat|] ==> eps(DD,ee,m,n): cont(DD`m,DD`n)";
-by (rtac (le_cases RS disjE) 1);
+by (rtac nat_linear_le 1);
by (resolve_tac prems 1);
by (rtac (hd(rev prems)) 1);
-by (asm_simp_tac(ZF_ss addsimps(eps_e_less::e_less_cont::prems)) 1);
-by (asm_simp_tac(ZF_ss addsimps(eps_e_gr::e_gr_cont::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(eps_e_less::e_less_cont::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(eps_e_gr::e_gr_cont::prems)) 1);
val eps_cont = result();
(* Theorems about splitting. *)
@@ -2123,7 +2097,7 @@
val prems = goal Limit.thy (* eps_split_add_left *)
"[|n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \
\ eps(DD,ee,m,m#+k) = eps(DD,ee,m#+n,m#+k) O eps(DD,ee,m,m#+n)";
-by (asm_simp_tac(arith_ss addsimps
+by (asm_simp_tac(!simpset addsimps
(eps_e_less::add_le_self::add_le_mono::prems)) 1);
brr(e_less_split_add::prems) 1;
val eps_split_add_left = result();
@@ -2131,7 +2105,7 @@
val prems = goal Limit.thy (* eps_split_add_left_rev *)
"[|n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \
\ eps(DD,ee,m,m#+n) = eps(DD,ee,m#+k,m#+n) O eps(DD,ee,m,m#+k)";
-by (asm_simp_tac(arith_ss addsimps
+by (asm_simp_tac(!simpset addsimps
(eps_e_less_add::eps_e_gr::add_le_self::add_le_mono::prems)) 1);
brr(e_less_e_gr_split_add::prems) 1;
val eps_split_add_left_rev = result();
@@ -2139,7 +2113,7 @@
val prems = goal Limit.thy (* eps_split_add_right *)
"[|m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \
\ eps(DD,ee,n#+k,n) = eps(DD,ee,n#+m,n) O eps(DD,ee,n#+k,n#+m)";
-by (asm_simp_tac(arith_ss addsimps
+by (asm_simp_tac(!simpset addsimps
(eps_e_gr::add_le_self::add_le_mono::prems)) 1);
brr(e_gr_split_add::prems) 1;
val eps_split_add_right = result();
@@ -2147,33 +2121,13 @@
val prems = goal Limit.thy (* eps_split_add_right_rev *)
"[|m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \
\ eps(DD,ee,n#+m,n) = eps(DD,ee,n#+k,n) O eps(DD,ee,n#+m,n#+k)";
-by (asm_simp_tac(arith_ss addsimps
+by (asm_simp_tac(!simpset addsimps
(eps_e_gr_add::eps_e_less::add_le_self::add_le_mono::prems)) 1);
brr(e_gr_e_less_split_add::prems) 1;
val eps_split_add_right_rev = result();
(* Arithmetic, little support in Isabelle/ZF. *)
-val prems = goal Arith.thy (* add_le_elim1 *)
- "[|m#+n le m#+k; m:nat; n:nat; k:nat|] ==> n le k";
-by (rtac mp 1);
-by (resolve_tac prems 2);
-by (res_inst_tac[("n","n")]nat_induct 1);
-by (resolve_tac prems 1);
-by (simp_tac (arith_ss addsimps prems) 1);
-by (safe_tac lemmas_cs);
-by (asm_full_simp_tac (ZF_ss addsimps
- (not_le_iff_lt::succ_le_iff::add_succ::add_succ_right::
- add_type::nat_into_Ord::prems)) 1);
-by (etac lt_asym 1);
-by (assume_tac 1);
-by (asm_full_simp_tac (ZF_ss addsimps add_succ_right::succ_le_iff::prems) 1);
-by (asm_full_simp_tac (ZF_ss addsimps
- (add_succ::le_iff::add_type::nat_into_Ord::prems)) 1);
-by (safe_tac lemmas_cs);
-by (etac lt_irrefl 1);
-val add_le_elim1 = result();
-
val prems = goal Limit.thy (* le_exists_lemma *)
"[|n le k; k le m; \
\ !!p q. [|p le q; k=n#+p; m=n#+q; p:nat; q:nat|] ==> R; \
@@ -2188,7 +2142,7 @@
by (assume_tac 2);
by (assume_tac 2);
by (cut_facts_tac[hd prems,hd(tl prems)]1);
-by (asm_full_simp_tac arith_ss 1);
+by (Asm_full_simp_tac 1);
by (etac add_le_elim1 1);
brr prems 1;
val le_exists_lemma = result();
@@ -2198,7 +2152,7 @@
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)";
by (rtac le_exists_lemma 1);
brr prems 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(eps_split_add_left::prems) 1;
val eps_split_left_le = result();
@@ -2207,7 +2161,7 @@
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)";
by (rtac le_exists_lemma 1);
brr prems 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(eps_split_add_left_rev::prems) 1;
val eps_split_left_le_rev = result();
@@ -2216,7 +2170,7 @@
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)";
by (rtac le_exists_lemma 1);
brr prems 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(eps_split_add_right::prems) 1;
val eps_split_right_le = result();
@@ -2225,7 +2179,7 @@
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)";
by (rtac le_exists_lemma 1);
brr prems 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(eps_split_add_right_rev::prems) 1;
val eps_split_right_le_rev = result();
@@ -2234,10 +2188,10 @@
val prems = goal Limit.thy (* eps_split_left *)
"[|m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)";
-by (rtac (le_cases RS disjE) 1);
+by (rtac nat_linear_le 1);
by (rtac eps_split_right_le_rev 4);
by (assume_tac 4);
-by (rtac (le_cases RS disjE) 3);
+by (rtac nat_linear_le 3);
by (rtac eps_split_left_le 5);
by (assume_tac 6);
by (rtac eps_split_left_le_rev 10);
@@ -2247,10 +2201,10 @@
val prems = goal Limit.thy (* eps_split_right *)
"[|n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)";
-by (rtac (le_cases RS disjE) 1);
+by (rtac nat_linear_le 1);
by (rtac eps_split_left_le_rev 3);
by (assume_tac 3);
-by (rtac (le_cases RS disjE) 8);
+by (rtac nat_linear_le 8);
by (rtac eps_split_right_le 10);
by (assume_tac 11);
by (rtac eps_split_right_le_rev 15);
@@ -2267,8 +2221,8 @@
"[|emb_chain(DD,ee); n:nat|] ==> \
\ rho_emb(DD,ee,n): set(DD`n) -> set(Dinf(DD,ee))";
brr(lam_type::DinfI::(eps_cont RS cont_fun RS apply_type)::prems) 1;
-by (asm_simp_tac arith_ss 1);
-by (rtac (le_cases RS disjE) 1);
+by (Asm_simp_tac 1);
+by (rtac nat_linear_le 1);
by (rtac nat_succI 1);
by (assume_tac 1);
by (resolve_tac prems 1);
@@ -2276,38 +2230,38 @@
but since x le y is x<succ(y) simplification does too much with this thm. *)
by (stac eps_split_right_le 1);
by (assume_tac 2);
-by (asm_simp_tac(ZF_ss addsimps (add1::[])) 1);
+by (asm_simp_tac(ZF_ss addsimps [add1]) 1);
brr(add_le_self::nat_0I::nat_succI::prems) 1;
-by (asm_simp_tac(ZF_ss addsimps(eps_succ_Rp::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(eps_succ_Rp::prems)) 1);
by (stac comp_fun_apply 1);
brr(eps_fun::nat_succI::(Rp_cont RS cont_fun)::emb_chain_emb::
emb_chain_cpo::refl::prems) 1;
(* Now the second part of the proof. Slightly different than HOL. *)
-by (asm_simp_tac(ZF_ss addsimps(eps_e_less::nat_succI::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(eps_e_less::nat_succI::prems)) 1);
by (etac (le_iff RS iffD1 RS disjE) 1);
-by (asm_simp_tac(ZF_ss addsimps(e_less_le::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(e_less_le::prems)) 1);
by (stac comp_fun_apply 1);
brr(e_less_cont::cont_fun::emb_chain_emb::emb_cont::prems) 1;
by (stac embRp_eq_thm 1);
brr(emb_chain_emb::(e_less_cont RS cont_fun RS apply_type)::emb_chain_cpo::
nat_succI::prems) 1;
-by (asm_simp_tac(ZF_ss addsimps(eps_e_less::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(eps_e_less::prems)) 1);
by (dtac shift_asm 1);
-by (asm_full_simp_tac(ZF_ss addsimps(eps_succ_Rp::e_less_eq::id_apply::
+by (asm_full_simp_tac(!simpset addsimps(eps_succ_Rp::e_less_eq::id_apply::
nat_succI::prems)) 1);
val rho_emb_fun = result();
val rho_emb_apply1 = prove_goalw Limit.thy [rho_emb_def]
"!!z. x:set(DD`n) ==> rho_emb(DD,ee,n)`x = (lam m:nat. eps(DD,ee,n,m)`x)"
- (fn prems => [asm_simp_tac ZF_ss 1]);
+ (fn prems => [Asm_simp_tac 1]);
val rho_emb_apply2 = prove_goalw Limit.thy [rho_emb_def]
"!!z. [|x:set(DD`n); m:nat|] ==> rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x"
- (fn prems => [asm_simp_tac ZF_ss 1]);
+ (fn prems => [Asm_simp_tac 1]);
val rho_emb_id = prove_goal Limit.thy
"!!z. [| x:set(DD`n); n:nat|] ==> rho_emb(DD,ee,n)`x`n = x"
- (fn prems => [asm_simp_tac(ZF_ss addsimps[rho_emb_apply2,eps_id,id_thm]) 1]);
+ (fn prems => [asm_simp_tac(!simpset addsimps[rho_emb_apply2,eps_id,id_thm]) 1]);
(* Shorter proof, 23 against 62. *)
@@ -2318,27 +2272,27 @@
brr(rho_emb_fun::prems) 1;
by (rtac rel_DinfI 1);
by (SELECT_GOAL(rewtac rho_emb_def) 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr((eps_cont RS cont_mono)::Dinf_prod::apply_type::rho_emb_fun::prems) 1;
(* Continuity, different order, slightly different proofs. *)
by (stac lub_Dinf 1);
by (rtac chainI 1);
brr(lam_type::(rho_emb_fun RS apply_type)::chain_in::prems) 1;
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (rtac rel_DinfI 1);
-by (asm_simp_tac(arith_ss addsimps (rho_emb_apply2::chain_in::[])) 1);
+by (asm_simp_tac(!simpset addsimps (rho_emb_apply2::chain_in::[])) 1);
brr((eps_cont RS cont_mono)::chain_rel::Dinf_prod::
(rho_emb_fun RS apply_type)::chain_in::nat_succI::prems) 1;
(* Now, back to the result of applying lub_Dinf *)
-by (asm_simp_tac(arith_ss addsimps (rho_emb_apply2::chain_in::[])) 1);
+by (asm_simp_tac(!simpset addsimps (rho_emb_apply2::chain_in::[])) 1);
by (stac rho_emb_apply1 1);
brr((cpo_lub RS islub_in)::emb_chain_cpo::prems) 1;
by (rtac fun_extension 1);
brr(lam_type::(eps_cont RS cont_fun RS apply_type)::(cpo_lub RS islub_in)::
emb_chain_cpo::prems) 1;
brr(cont_chain::eps_cont::emb_chain_cpo::prems) 1;
-by (asm_simp_tac arith_ss 1);
-by (asm_simp_tac(ZF_ss addsimps((eps_cont RS cont_lub)::prems)) 1);
+by (Asm_simp_tac 1);
+by (asm_simp_tac(!simpset addsimps((eps_cont RS cont_lub)::prems)) 1);
val rho_emb_cont = result();
(* 32 vs 61, using safe_tac with imp in asm would be unfortunate (5steps) *)
@@ -2349,10 +2303,10 @@
by (rtac impE 1 THEN atac 3 THEN rtac(hd prems) 2); (* For induction proof *)
by (res_inst_tac[("n","n")]nat_induct 1);
by (rtac impI 2);
-by (asm_full_simp_tac (arith_ss addsimps (e_less_eq::prems)) 2);
+by (asm_full_simp_tac (!simpset addsimps (e_less_eq::prems)) 2);
by (stac id_thm 2);
brr(apply_type::Dinf_prod::cpo_refl::emb_chain_cpo::nat_0I::prems) 1;
-by (asm_full_simp_tac (arith_ss addsimps [le_succ_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [le_succ_iff]) 1);
by (rtac impI 1);
by (etac disjE 1);
by (dtac mp 1 THEN atac 1);
@@ -2377,7 +2331,7 @@
brr((hd(tl(tl prems)) RS Dinf_prod RS apply_type)::cont_fun::Rp_cont::
e_less_cont::emb_cont::emb_chain_emb::emb_chain_cpo::apply_type::
embRp_rel::(disjI1 RS (le_succ_iff RS iffD2))::nat_succI::prems) 1;
-by (asm_full_simp_tac (arith_ss addsimps (e_less_eq::prems)) 1);
+by (asm_full_simp_tac (!simpset addsimps (e_less_eq::prems)) 1);
by (stac id_thm 1);
brr(apply_type::Dinf_prod::cpo_refl::emb_chain_cpo::nat_succI::prems) 1;
val lemma1 = result();
@@ -2390,10 +2344,10 @@
by (rtac impE 1 THEN atac 3 THEN rtac(hd prems) 2); (* For induction proof *)
by (res_inst_tac[("n","m")]nat_induct 1);
by (rtac impI 2);
-by (asm_full_simp_tac (arith_ss addsimps (e_gr_eq::prems)) 2);
+by (asm_full_simp_tac (!simpset addsimps (e_gr_eq::prems)) 2);
by (stac id_thm 2);
brr(apply_type::Dinf_prod::cpo_refl::emb_chain_cpo::nat_0I::prems) 1;
-by (asm_full_simp_tac (arith_ss addsimps [le_succ_iff]) 1);
+by (asm_full_simp_tac (!simpset addsimps [le_succ_iff]) 1);
by (rtac impI 1);
by (etac disjE 1);
by (dtac mp 1 THEN atac 1);
@@ -2402,24 +2356,17 @@
by (stac Dinf_eq 7);
brr(emb_chain_emb::emb_chain_cpo::Rp_cont::e_gr_cont::cont_fun::emb_cont::
apply_type::Dinf_prod::nat_succI::prems) 1;
-by (asm_full_simp_tac (arith_ss addsimps (e_gr_eq::prems)) 1);
+by (asm_full_simp_tac (!simpset addsimps (e_gr_eq::prems)) 1);
by (stac id_thm 1);
brr(apply_type::Dinf_prod::cpo_refl::emb_chain_cpo::nat_succI::prems) 1;
val lemma2 = result();
-val prems = goalw ZF.thy [if_def]
- "[| P==>R(a); ~P==>R(b) |] ==> R(if(P,a,b))";
-by (excluded_middle_tac"P"1);
-by (ALLGOALS(asm_simp_tac ZF_ss THEN' rtac theI2));
-by (ALLGOALS(asm_full_simp_tac FOL_ss));
-brr(ex1I::refl::prems) 1;
-val if_case = result();
-
val prems = goalw Limit.thy [eps_def] (* eps1 *)
"[|emb_chain(DD,ee); x:set(Dinf(DD,ee)); m:nat; n:nat|] ==> \
\ rel(DD`n,eps(DD,ee,m,n)`(x`m),x`n)";
-by (rtac if_case 1);
-brr(lemma1::(not_le_iff_lt RS iffD1 RS leI RS lemma2)::nat_into_Ord::prems) 1;
+by (split_tac [expand_if] 1);
+brr(conjI::impI::lemma1::
+ (not_le_iff_lt RS iffD1 RS leI RS lemma2)::nat_into_Ord::prems) 1;
val eps1 = result();
(* The following theorem is needed/useful due to type check for rel_cfI,
@@ -2431,11 +2378,11 @@
\ (lam x:set(Dinf(DD,ee)). x`n) : cont(Dinf(DD,ee),DD`n)";
by (rtac contI 1);
brr(lam_type::apply_type::Dinf_prod::prems) 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(rel_Dinf::prems) 1;
by (stac beta 1);
brr(cpo_Dinf::islub_in::cpo_lub::prems) 1;
-by (asm_simp_tac(ZF_ss addsimps(chain_in::lub_Dinf::prems)) 1);
+by (asm_simp_tac(!simpset addsimps(chain_in::lub_Dinf::prems)) 1);
val lam_Dinf_cont = result();
val prems = goalw Limit.thy [rho_proj_def] (* rho_projpair *)
@@ -2487,19 +2434,19 @@
"[| !!n. n:nat ==> emb(DD`n,E,r(n)); \
\ !!m n. [|m le n; m:nat; n:nat|] ==> r(n) O eps(DD,ee,m,n) = r(m) |] ==> \
\ commute(DD,ee,E,r)";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
brr prems 1;
val commuteI = result();
val prems = goalw Limit.thy [commute_def] (* commute_emb *)
"!!z. [| commute(DD,ee,E,r); n:nat |] ==> emb(DD`n,E,r(n))";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val commute_emb = result();
val prems = goalw Limit.thy [commute_def] (* commute_eq *)
"!!z. [| commute(DD,ee,E,r); m le n; m:nat; n:nat |] ==> \
\ r(n) O eps(DD,ee,m,n) = r(m) ";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
val commute_eq = result();
(* Shorter proof: 11 vs 46 lines. *)
@@ -2513,7 +2460,7 @@
by (rtac fun_extension 6); (* Next, clean up and instantiate unknowns *)
brr(comp_fun::rho_emb_fun::eps_fun::Dinf_prod::apply_type::prems) 1;
by (asm_simp_tac
- (ZF_ss addsimps(rho_emb_apply2::(eps_fun RS apply_type)::prems)) 1);
+ (!simpset addsimps(rho_emb_apply2::(eps_fun RS apply_type)::prems)) 1);
by (rtac (comp_fun_apply RS subst) 1);
by (rtac (eps_split_left RS subst) 4);
brr(eps_fun::refl::prems) 1;
@@ -2533,7 +2480,7 @@
by (rtac chainI 1);
brr(lam_type::cont_cf::comp_pres_cont::emb_r::Rp_cont::emb_cont::
emb_chain_cpo::prems) 1;
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (res_inst_tac[("r1","r"),("m1","n")](commute_eq RS subst) 1);
brr(le_succ::nat_succI::prems) 1;
by (stac Rp_comp 1);
@@ -2565,7 +2512,7 @@
\ lam n:nat. \
\ (rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))`x)";
by (cut_facts_tac[hd(tl prems) RS (hd prems RS (rho_emb_chain RS chain_cf))]1);
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val rho_emb_chain_apply1 = result();
val prems = goal Limit.thy
@@ -2584,7 +2531,7 @@
by (cut_facts_tac
[hd(tl(tl prems)) RS (hd prems RS (hd(tl prems) RS (hd prems RS
(rho_emb_chain_apply1 RS chain_Dinf RS chain_iprod_emb_chain))))]1);
-by (asm_full_simp_tac ZF_ss 1);
+by (Asm_full_simp_tac 1);
val rho_emb_chain_apply2 = result();
(* Shorter proof: 32 vs 72 (roughly), Isabelle proof has lemmas. *)
@@ -2599,17 +2546,17 @@
brr(cpo_Dinf::prems) 1;
by (rtac islub_least 1);
brr(cpo_lub::rho_emb_chain::cpo_cf::cpo_Dinf::isubI::cont_cf::id_cont::prems) 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(embRp_rel::emb_rho_emb::emb_chain_cpo::cpo_Dinf::prems) 1;
by (rtac rel_cfI 1);
by (asm_simp_tac
- (ZF_ss addsimps(id_thm::lub_cf::rho_emb_chain::cpo_Dinf::prems)) 1);
+ (!simpset addsimps(id_thm::lub_cf::rho_emb_chain::cpo_Dinf::prems)) 1);
by (rtac rel_DinfI 1); (* Addtional assumptions *)
by (stac lub_Dinf 1);
brr(rho_emb_chain_apply1::prems) 1;
brr(Dinf_prod::(cpo_lub RS islub_in)::id_cont::cpo_Dinf::cpo_cf::cf_cont::
rho_emb_chain::rho_emb_chain_apply1::(id_cont RS cont_cf)::prems) 2;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (rtac dominate_islub 1);
by (rtac cpo_lub 3);
brr(rho_emb_chain_apply2::emb_chain_cpo::prems) 3;
@@ -2618,12 +2565,12 @@
rho_emb_chain_apply2::prems) 2;
by (rtac dominateI 1);
by (assume_tac 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (stac comp_fun_apply 1);
brr(cont_fun::Rp_cont::emb_cont::emb_rho_emb::cpo_Dinf::emb_chain_cpo::prems) 1;
by (stac ((rho_projpair RS Rp_unique)) 1);
by (SELECT_GOAL(rewtac rho_proj_def) 5);
-by (asm_simp_tac ZF_ss 5);
+by (Asm_simp_tac 5);
by (stac rho_emb_id 5);
brr(cpo_refl::cpo_Dinf::apply_type::Dinf_prod::emb_chain_cpo::prems) 1;
val rho_emb_lub = result();
@@ -2637,7 +2584,7 @@
by (rtac chainI 1);
brr(lam_type::cont_cf::comp_pres_cont::emb_r::emb_f::
Rp_cont::emb_cont::emb_chain_cpo::prems) 1;
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (res_inst_tac[("r1","r"),("m1","n")](commute_eq RS subst) 1);
by (res_inst_tac[("r1","f"),("m1","n")](commute_eq RS subst) 5);
brr(le_succ::nat_succI::prems) 1;
@@ -2664,7 +2611,7 @@
by (rtac chainI 1);
brr(lam_type::cont_cf::comp_pres_cont::emb_r::emb_f::
Rp_cont::emb_cont::emb_chain_cpo::prems) 1;
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
by (res_inst_tac[("r1","r"),("m1","n")](commute_eq RS subst) 1);
by (res_inst_tac[("r1","f"),("m1","n")](commute_eq RS subst) 5);
brr(le_succ::nat_succI::prems) 1;
@@ -2703,7 +2650,7 @@
val lemma = result();
val lemma_assoc = prove_goal Limit.thy "a O b O c O d = a O (b O c) O d"
- (fn prems => [simp_tac (ZF_ss addsimps[comp_assoc]) 1]);
+ (fn prems => [simp_tac (!simpset addsimps[comp_assoc]) 1]);
fun elem n l = if n = 1 then hd l else elem(n-1)(tl l);
@@ -2717,16 +2664,16 @@
\ (E,G, \
\ lub(cf(E,G), lam n:nat. f(n) O Rp(DD`n,E,r(n))), \
\ lub(cf(G,E), lam n:nat. r(n) O Rp(DD`n,G,f(n))))";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
by (stac comp_lubs 3);
(* The following one line is 15 lines in HOL, and includes existentials. *)
brr(cf_cont::islub_in::cpo_lub::cpo_cf::theta_chain::theta_proj_chain::prems) 1;
-by (simp_tac (ZF_ss addsimps[comp_assoc]) 1);
-by (simp_tac (ZF_ss addsimps[(tl prems) MRS lemma]) 1);
+by (simp_tac (!simpset addsimps[comp_assoc]) 1);
+by (simp_tac (!simpset addsimps[(tl prems) MRS lemma]) 1);
by (stac comp_lubs 2);
brr(cf_cont::islub_in::cpo_lub::cpo_cf::theta_chain::theta_proj_chain::prems) 1;
-by (simp_tac (ZF_ss addsimps[comp_assoc]) 1);
-by (simp_tac (ZF_ss addsimps[
+by (simp_tac (!simpset addsimps[comp_assoc]) 1);
+by (simp_tac (!simpset addsimps[
[elem 3 prems,elem 2 prems,elem 4 prems,elem 6 prems, elem 5 prems]
MRS lemma]) 1);
by (rtac dominate_islub 1);
@@ -2735,7 +2682,7 @@
chain_fun::chain_const::prems) 2;
by (rtac dominateI 1);
by (assume_tac 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(embRp_rel::emb_f::emb_chain_cpo::prems) 1;
val theta_projpair = result();
@@ -2753,10 +2700,10 @@
\ (lam f : cont(D',E). f O g) : mono(cf(D',E),cf(D,E))";
by (rtac monoI 1);
by (REPEAT(dtac cf_cont 2));
-by (asm_simp_tac ZF_ss 2);
+by (Asm_simp_tac 2);
by (rtac comp_mono 2);
by (SELECT_GOAL(rewrite_goals_tac[set_def,cf_def]) 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
brr(lam_type::comp_pres_cont::cpo_cf::cpo_refl::cont_cf::prems) 1;
val mono_lemma = result();
@@ -2775,7 +2722,7 @@
by (stac beta 5);
by (rtac lam_type 1);
by (stac beta 1);
-by (ALLGOALS(asm_simp_tac (ZF_ss addsimps prems)));
+by (ALLGOALS(asm_simp_tac (!simpset addsimps prems)));
brr(lam_type::comp_pres_cont::Rp_cont::emb_cont::emb_r::emb_f::
emb_chain_cpo::prems) 1;
val lemma = result();
@@ -2796,11 +2743,11 @@
"[| commute(DD,ee,E,r); commute(DD,ee,G,f); \
\ emb_chain(DD,ee); cpo(E); cpo(G); cpo(DD`x); x:nat |] ==> \
\ suffix(lam n:nat. (f(n) O Rp(DD`n,E,r(n))) O r(x),x) = (lam n:nat. f(x))";
-by (simp_tac (arith_ss addsimps prems) 1);
+by (simp_tac (!simpset addsimps prems) 1);
by (rtac fun_extension 1);
brr(lam_type::comp_fun::cont_fun::Rp_cont::emb_cont::emb_r::emb_f::
add_type::emb_chain_cpo::prems) 1;
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (res_inst_tac[("r1","r"),("m1","x")](commute_eq RS subst) 1);
brr(emb_r::add_le_self::add_type::prems) 1;
by (stac comp_assoc 1);
@@ -2812,16 +2759,16 @@
val suffix_lemma = result();
val mediatingI = prove_goalw Limit.thy [mediating_def]
- "[|emb(E,G,t);!!n.n:nat ==> f(n) = t O r(n) |]==>mediating(E,G,r,f,t)"
- (fn prems => [safe_tac lemmas_cs,trr prems 1]);
+ "[|emb(E,G,t); !!n.n:nat ==> f(n) = t O r(n) |]==>mediating(E,G,r,f,t)"
+ (fn prems => [safe_tac (!claset),trr prems 1]);
val mediating_emb = prove_goalw Limit.thy [mediating_def]
"!!z. mediating(E,G,r,f,t) ==> emb(E,G,t)"
- (fn prems => [fast_tac ZF_cs 1]);
+ (fn prems => [Fast_tac 1]);
val mediating_eq = prove_goalw Limit.thy [mediating_def]
"!!z. [| mediating(E,G,r,f,t); n:nat |] ==> f(n) = t O r(n)"
- (fn prems => [fast_tac ZF_cs 1]);
+ (fn prems => [Fast_tac 1]);
val prems = goal Limit.thy (* lub_universal_mediating *)
"[| lub(cf(E,E), lam n:nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); \
@@ -2833,7 +2780,7 @@
by (stac comp_lubs 3);
brr(cont_cf::emb_cont::emb_r::cpo_cf::theta_chain::chain_const::
emb_chain_cpo::prems) 1;
-by (simp_tac ZF_ss 1);
+by (Simp_tac 1);
by (rtac (lub_suffix RS subst) 1);
brr(chain_lemma::cpo_cf::emb_chain_cpo::prems) 1;
by (stac (tl prems MRS suffix_lemma) 1);
@@ -2851,7 +2798,7 @@
by (rtac (hd(tl prems) RS subst) 2);
by (res_inst_tac[("b","t")](lub_const RS subst) 2);
by (stac comp_lubs 4);
-by (simp_tac (ZF_ss addsimps(comp_assoc::(hd prems RS mediating_eq)::prems)) 9);
+by (simp_tac (!simpset addsimps(comp_assoc::(hd prems RS mediating_eq)::prems)) 9);
brr(cont_fun::emb_cont::mediating_emb::cont_cf::cpo_cf::chain_const::
commute_chain::emb_chain_cpo::prems) 1;
val lub_universal_unique = result();
@@ -2870,7 +2817,7 @@
\ (ALL t. mediating(Dinf(DD,ee),G,rho_emb(DD,ee),f,t) --> \
\ t = lub(cf(Dinf(DD,ee),G), \
\ lam n:nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))))";
-by (safe_tac lemmas_cs);
+by (safe_tac (!claset));
brr(lub_universal_mediating::rho_emb_commute::rho_emb_lub::cpo_Dinf::prems) 1;
brr(lub_universal_unique::rho_emb_commute::rho_emb_lub::cpo_Dinf::prems) 1;
val Dinf_universal = result();
--- a/src/ZF/ex/Limit.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Limit.thy Fri Jan 03 15:01:55 1997 +0100
@@ -17,7 +17,7 @@
Laboratory, 1995.
*)
-Limit = ZF + Perm + Arith +
+Limit = Perm + Arith +
consts
--- a/src/ZF/ex/ListN.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/ListN.ML Fri Jan 03 15:01:55 1997 +0100
@@ -13,20 +13,20 @@
goal ListN.thy "!!l. l:list(A) ==> <length(l),l> : listn(A)";
by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac list_ss));
+by (ALLGOALS Asm_simp_tac);
by (REPEAT (ares_tac listn.intrs 1));
qed "list_into_listn";
goal ListN.thy "<n,l> : listn(A) <-> l:list(A) & length(l)=n";
by (rtac iffI 1);
by (etac listn.induct 1);
-by (safe_tac (ZF_cs addSIs (list_typechecks @
+by (safe_tac (!claset addSIs (list_typechecks @
[length_Nil, length_Cons, list_into_listn])));
qed "listn_iff";
goal ListN.thy "listn(A)``{n} = {l:list(A). length(l)=n}";
by (rtac equality_iffI 1);
-by (simp_tac (list_ss addsimps [listn_iff,separation,image_singleton_iff]) 1);
+by (simp_tac (!simpset addsimps [listn_iff,separation,image_singleton_iff]) 1);
qed "listn_image_eq";
goalw ListN.thy listn.defs "!!A B. A<=B ==> listn(A) <= listn(B)";
@@ -39,7 +39,7 @@
"!!n l. [| <n,l> : listn(A); <n',l'> : listn(A) |] ==> \
\ <n#+n', l@l'> : listn(A)";
by (etac listn.induct 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps listn.intrs)));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps listn.intrs)));
qed "listn_append";
val Nil_listn_case = listn.mk_cases list.con_defs "<i,Nil> : listn(A)"
--- a/src/ZF/ex/Mutil.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Mutil.ML Fri Jan 03 15:01:55 1997 +0100
@@ -16,36 +16,37 @@
qed "evnodd_iff";
goalw thy [evnodd_def] "evnodd(A, b) <= A";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "evnodd_subset";
(* Finite(X) ==> Finite(evnodd(X,b)) *)
bind_thm("Finite_evnodd", evnodd_subset RS subset_imp_lepoll RS lepoll_Finite);
goalw thy [evnodd_def] "evnodd(A Un B, b) = evnodd(A,b) Un evnodd(B,b)";
-by (simp_tac (ZF_ss addsimps [Collect_Un]) 1);
+by (simp_tac (!simpset addsimps [Collect_Un]) 1);
qed "evnodd_Un";
goalw thy [evnodd_def] "evnodd(A - B, b) = evnodd(A,b) - evnodd(B,b)";
-by (simp_tac (ZF_ss addsimps [Collect_Diff]) 1);
+by (simp_tac (!simpset addsimps [Collect_Diff]) 1);
qed "evnodd_Diff";
goalw thy [evnodd_def]
"evnodd(cons(<i,j>,C), b) = \
\ if((i#+j) mod 2 = b, cons(<i,j>, evnodd(C,b)), evnodd(C,b))";
-by (asm_simp_tac (ZF_ss addsimps [evnodd_def, Collect_cons]
+by (asm_simp_tac (!simpset addsimps [evnodd_def, Collect_cons]
setloop split_tac [expand_if]) 1);
qed "evnodd_cons";
goalw thy [evnodd_def] "evnodd(0, b) = 0";
-by (simp_tac (ZF_ss addsimps [evnodd_def]) 1);
+by (simp_tac (!simpset addsimps [evnodd_def]) 1);
qed "evnodd_0";
+Addsimps [evnodd_cons, evnodd_0];
(*** Dominoes ***)
goal thy "!!d. d:domino ==> Finite(d)";
-by (fast_tac (ZF_cs addSIs [Finite_cons, Finite_0] addEs [domino.elim]) 1);
+by (fast_tac (!claset addSIs [Finite_cons, Finite_0] addEs [domino.elim]) 1);
qed "domino_Finite";
goal thy "!!d. [| d:domino; b<2 |] ==> EX i' j'. evnodd(d,b) = {<i',j'>}";
@@ -54,10 +55,9 @@
by (res_inst_tac [("k1", "i#+j")] (mod2_cases RS disjE) 1);
by (REPEAT_FIRST (ares_tac [add_type]));
(*Four similar cases: case (i#+j) mod 2 = b, 2#-b, ...*)
-by (REPEAT (asm_simp_tac (arith_ss addsimps [evnodd_cons, evnodd_0, mod_succ,
- succ_neq_self]
+by (REPEAT (asm_simp_tac (!simpset addsimps [mod_succ, succ_neq_self]
setloop split_tac [expand_if]) 1
- THEN fast_tac (ZF_cs addDs [ltD]) 1));
+ THEN fast_tac (!claset addDs [ltD]) 1));
qed "domino_singleton";
@@ -68,50 +68,50 @@
goal thy "!!t. t: tiling(A) ==> \
\ u: tiling(A) --> t Int u = 0 --> t Un u : tiling(A)";
by (etac tiling.induct 1);
-by (simp_tac (ZF_ss addsimps tiling.intrs) 1);
-by (fast_tac (ZF_cs addIs tiling.intrs
- addss (ZF_ss addsimps [Un_assoc,
- subset_empty_iff RS iff_sym])) 1);
+by (simp_tac (!simpset addsimps tiling.intrs) 1);
+by (fast_tac (!claset addIs tiling.intrs
+ addss (!simpset addsimps [Un_assoc,
+ subset_empty_iff RS iff_sym])) 1);
bind_thm ("tiling_UnI", result() RS mp RS mp);
goal thy "!!t. t:tiling(domino) ==> Finite(t)";
by (eresolve_tac [tiling.induct] 1);
by (resolve_tac [Finite_0] 1);
-by (fast_tac (ZF_cs addIs [domino_Finite, Finite_Un]) 1);
+by (fast_tac (!claset addIs [domino_Finite, Finite_Un]) 1);
qed "tiling_domino_Finite";
goal thy "!!t. t: tiling(domino) ==> |evnodd(t,0)| = |evnodd(t,1)|";
by (eresolve_tac [tiling.induct] 1);
-by (simp_tac (ZF_ss addsimps [evnodd_def]) 1);
+by (simp_tac (!simpset addsimps [evnodd_def]) 1);
by (res_inst_tac [("b1","0")] (domino_singleton RS exE) 1);
-by (simp_tac arith_ss 2 THEN assume_tac 1);
+by (Simp_tac 2 THEN assume_tac 1);
by (res_inst_tac [("b1","1")] (domino_singleton RS exE) 1);
-by (simp_tac arith_ss 2 THEN assume_tac 1);
-by (step_tac ZF_cs 1);
+by (Simp_tac 2 THEN assume_tac 1);
+by (step_tac (!claset) 1);
by (subgoal_tac "ALL p b. p:evnodd(a,b) --> p~:evnodd(ta,b)" 1);
-by (asm_simp_tac (ZF_ss addsimps [evnodd_Un, Un_cons, tiling_domino_Finite,
- evnodd_subset RS subset_Finite,
- Finite_imp_cardinal_cons]) 1);
-by (fast_tac (ZF_cs addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1);
+by (asm_simp_tac (!simpset addsimps [evnodd_Un, Un_cons, tiling_domino_Finite,
+ evnodd_subset RS subset_Finite,
+ Finite_imp_cardinal_cons]) 1);
+by (fast_tac (!claset addSDs [evnodd_subset RS subsetD] addEs [equalityE]) 1);
qed "tiling_domino_0_1";
goal thy "!!i n. [| i: nat; n: nat |] ==> {i} * (n #+ n) : tiling(domino)";
by (nat_ind_tac "n" [] 1);
-by (simp_tac (arith_ss addsimps tiling.intrs) 1);
-by (asm_simp_tac (arith_ss addsimps [Un_assoc RS sym, Sigma_succ2]) 1);
+by (simp_tac (!simpset addsimps tiling.intrs) 1);
+by (asm_simp_tac (!simpset addsimps [Un_assoc RS sym, Sigma_succ2]) 1);
by (resolve_tac tiling.intrs 1);
by (assume_tac 2);
-by (subgoal_tac (*seems the easiest way of turning one to the other*)
+by (subgoal_tac (*seems the easiest way of turning one to the other*)
"{i}*{succ(n1#+n1)} Un {i}*{n1#+n1} = {<i,n1#+n1>, <i,succ(n1#+n1)>}" 1);
by (fast_tac eq_cs 2);
-by (asm_simp_tac (arith_ss addsimps [domino.horiz]) 1);
+by (asm_simp_tac (!simpset addsimps [domino.horiz]) 1);
by (fast_tac (eq_cs addEs [mem_irrefl, mem_asym]) 1);
qed "dominoes_tile_row";
goal thy "!!m n. [| m: nat; n: nat |] ==> m * (n #+ n) : tiling(domino)";
by (nat_ind_tac "m" [] 1);
-by (simp_tac (arith_ss addsimps tiling.intrs) 1);
-by (asm_simp_tac (arith_ss addsimps [Sigma_succ1]) 1);
+by (simp_tac (!simpset addsimps tiling.intrs) 1);
+by (asm_simp_tac (!simpset addsimps [Sigma_succ1]) 1);
by (fast_tac (eq_cs addIs [tiling_UnI, dominoes_tile_row]
addEs [mem_irrefl]) 1);
qed "dominoes_tile_matrix";
@@ -124,23 +124,23 @@
by (resolve_tac [notI] 1);
by (dresolve_tac [tiling_domino_0_1] 1);
by (subgoal_tac "|evnodd(t',0)| < |evnodd(t',1)|" 1);
-by (asm_full_simp_tac (ZF_ss addsimps [lt_not_refl]) 1);
+by (asm_full_simp_tac (!simpset addsimps [lt_not_refl]) 1);
by (subgoal_tac "t : tiling(domino)" 1);
(*Requires a small simpset that won't move the succ applications*)
by (asm_simp_tac (ZF_ss addsimps [nat_succI, add_type,
- dominoes_tile_matrix]) 2);
+ dominoes_tile_matrix]) 2);
by (subgoal_tac "(m#+m)#+(n#+n) = (m#+n)#+(m#+n)" 1);
-by (asm_simp_tac (arith_ss addsimps add_ac) 2);
+by (asm_simp_tac (!simpset addsimps add_ac) 2);
by (asm_full_simp_tac
- (arith_ss addsimps [evnodd_Diff, evnodd_cons, evnodd_0, mod2_add_self,
- mod2_succ_succ, tiling_domino_0_1 RS sym]) 1);
+ (!simpset addsimps [evnodd_Diff, mod2_add_self,
+ mod2_succ_succ, tiling_domino_0_1 RS sym]) 1);
by (resolve_tac [lt_trans] 1);
by (REPEAT
(rtac Finite_imp_cardinal_Diff 1
THEN
- asm_simp_tac (arith_ss addsimps [tiling_domino_Finite, Finite_evnodd,
- Finite_Diff]) 1
+ asm_simp_tac (!simpset addsimps [tiling_domino_Finite, Finite_evnodd,
+ Finite_Diff]) 1
THEN
- asm_simp_tac (arith_ss addsimps [evnodd_iff, nat_0_le RS ltD,
- mod2_add_self]) 1));
+ asm_simp_tac (!simpset addsimps [evnodd_iff, nat_0_le RS ltD,
+ mod2_add_self]) 1));
qed "mutil_not_tiling";
--- a/src/ZF/ex/Ntree.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Ntree.ML Fri Jan 03 15:01:55 1997 +0100
@@ -29,8 +29,8 @@
by (rtac (major RS ntree.induct) 1);
by (etac UN_E 1);
by (REPEAT_SOME (ares_tac prems));
-by (fast_tac (ZF_cs addEs [fun_weaken_type]) 1);
-by (fast_tac (ZF_cs addDs [apply_type]) 1);
+by (fast_tac (!claset addEs [fun_weaken_type]) 1);
+by (fast_tac (!claset addDs [apply_type]) 1);
qed "ntree_induct";
(*Induction on ntree(A) to prove an equation*)
@@ -44,7 +44,7 @@
by (cut_facts_tac prems 1);
by (rtac fun_extension 1);
by (REPEAT_SOME (ares_tac [comp_fun]));
-by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply]) 1);
+by (asm_simp_tac (!simpset addsimps [comp_fun_apply]) 1);
qed "ntree_induct_eqn";
(** Lemmas to justify using "Ntree" in other recursive type definitions **)
@@ -59,7 +59,7 @@
goalw Ntree.thy (ntree.defs@ntree.con_defs) "ntree(univ(A)) <= univ(A)";
by (rtac lfp_lowerbound 1);
by (rtac (A_subset_univ RS univ_mono) 2);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (REPEAT (ares_tac [Pair_in_univ, nat_fun_univ RS subsetD] 1));
qed "ntree_univ";
@@ -87,7 +87,7 @@
by (eresolve_tac [Collect_subset RS FiniteFun_mono1 RS subsetD] 1);
by (dresolve_tac [FiniteFun.dom_subset RS subsetD] 1);
by (dresolve_tac [Fin.dom_subset RS subsetD] 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "maptree_induct";
@@ -112,6 +112,6 @@
FiniteFun_mono RS subsetD] 1);
by (dresolve_tac [FiniteFun.dom_subset RS subsetD] 1);
by (dresolve_tac [Fin.dom_subset RS subsetD] 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "maptree2_induct";
--- a/src/ZF/ex/Primes.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Primes.ML Fri Jan 03 15:01:55 1997 +0100
@@ -23,24 +23,24 @@
goalw thy [dvd_def] "!!m. m:nat ==> m dvd 0";
-by (fast_tac (ZF_cs addIs [nat_0I, mult_0_right RS sym]) 1);
+by (fast_tac (!claset addIs [nat_0I, mult_0_right RS sym]) 1);
qed "dvd_0_right";
goalw thy [dvd_def] "!!m. 0 dvd m ==> m = 0";
-by (fast_tac (ZF_cs addss arith_ss) 1);
+by (fast_tac (!claset addss (!simpset)) 1);
qed "dvd_0_left";
goalw thy [dvd_def] "!!m. m:nat ==> m dvd m";
-by (fast_tac (ZF_cs addIs [nat_1I, mult_1_right RS sym]) 1);
+by (fast_tac (!claset addIs [nat_1I, mult_1_right RS sym]) 1);
qed "dvd_refl";
goalw thy [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p";
-by (fast_tac (ZF_cs addIs [mult_assoc, mult_type] ) 1);
+by (fast_tac (!claset addIs [mult_assoc, mult_type] ) 1);
qed "dvd_trans";
goalw thy [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n";
-by (fast_tac (ZF_cs addDs [mult_eq_self_implies_10]
- addss (arith_ss addsimps [mult_assoc, mult_eq_1_iff])) 1);
+by (fast_tac (!claset addDs [mult_eq_self_implies_10]
+ addss (!simpset addsimps [mult_assoc, mult_eq_1_iff])) 1);
qed "dvd_anti_sym";
@@ -52,35 +52,35 @@
goalw thy [egcd_def] "!!m. m:nat ==> egcd(m,0) = m";
by (stac transrec 1);
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
qed "egcd_0";
goalw thy [egcd_def]
"!!m. [| 0<n; m:nat; n:nat |] ==> egcd(m,n) = egcd(n, m mod n)";
by (res_inst_tac [("P", "%z. ?left(z) = ?right")] (transrec RS ssubst) 1);
-by (asm_simp_tac (arith_ss addsimps [ltD RS mem_imp_not_eq RS not_sym,
+by (asm_simp_tac (!simpset addsimps [ltD RS mem_imp_not_eq RS not_sym,
mod_less_divisor RS ltD]) 1);
qed "egcd_lt_0";
goal thy "!!m. m:nat ==> egcd(m,0) dvd m";
-by (asm_simp_tac (arith_ss addsimps [egcd_0,dvd_refl]) 1);
+by (asm_simp_tac (!simpset addsimps [egcd_0,dvd_refl]) 1);
qed "egcd_0_dvd_m";
goal thy "!!m. m:nat ==> egcd(m,0) dvd 0";
-by (asm_simp_tac (arith_ss addsimps [egcd_0,dvd_0_right]) 1);
+by (asm_simp_tac (!simpset addsimps [egcd_0,dvd_0_right]) 1);
qed "egcd_0_dvd_0";
goalw thy [dvd_def] "!!k. [| k dvd a; k dvd b |] ==> k dvd (a #+ b)";
-by (fast_tac (ZF_cs addIs [add_mult_distrib_left RS sym, add_type]) 1);
+by (fast_tac (!claset addIs [add_mult_distrib_left RS sym, add_type]) 1);
qed "dvd_add";
goalw thy [dvd_def] "!!k. [| k dvd a; q:nat |] ==> k dvd (q #* a)";
-by (fast_tac (ZF_cs addIs [mult_left_commute, mult_type]) 1);
+by (fast_tac (!claset addIs [mult_left_commute, mult_type]) 1);
qed "dvd_mult";
goal thy "!!k. [| k dvd b; k dvd (a mod b); 0 < b; a:nat |] ==> k dvd a";
by (deepen_tac
- (ZF_cs addIs [mod_div_equality RS subst]
+ (!claset addIs [mod_div_equality RS subst]
addDs [dvdD]
addSIs [dvd_add, dvd_mult, mult_type,mod_type,div_type]) 0 1);
qed "gcd_ind";
@@ -93,12 +93,12 @@
by (rtac ballI 1);
by (excluded_middle_tac "x=0" 1);
(* case x = 0 *)
-by (asm_simp_tac (arith_ss addsimps [egcd_0]) 2);
+by (asm_simp_tac (!simpset addsimps [egcd_0]) 2);
(* case x > 0 *)
-by (asm_simp_tac (arith_ss addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt]) 1);
+by (asm_simp_tac (!simpset addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt]) 1);
by (eres_inst_tac [("x","a mod x")] ballE 1);
-by (asm_simp_tac ZF_ss 1);
-by (asm_full_simp_tac (arith_ss addsimps [mod_less_divisor RS ltD,
+by (Asm_simp_tac 1);
+by (asm_full_simp_tac (!simpset addsimps [mod_less_divisor RS ltD,
nat_into_Ord RS Ord_0_lt]) 1);
qed "egcd_type";
@@ -111,31 +111,31 @@
by (rtac ballI 1);
by (excluded_middle_tac "x=0" 1);
(* case x = 0 *)
-by (asm_simp_tac (arith_ss addsimps [egcd_0,dvd_refl,dvd_0_right]) 2);
+by (asm_simp_tac (!simpset addsimps [egcd_0,dvd_refl,dvd_0_right]) 2);
(* case x > 0 *)
-by (asm_simp_tac (arith_ss addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt]) 1);
+by (asm_simp_tac (!simpset addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt]) 1);
by (eres_inst_tac [("x","a mod x")] ballE 1);
-by (asm_simp_tac ZF_ss 1);
-by (asm_full_simp_tac (arith_ss addsimps [mod_less_divisor RS ltD,
+by (Asm_simp_tac 1);
+by (asm_full_simp_tac (!simpset addsimps [mod_less_divisor RS ltD,
nat_into_Ord RS Ord_0_lt]) 2);
-by (best_tac (ZF_cs addIs [gcd_ind, nat_into_Ord RS Ord_0_lt]) 1);
+by (best_tac (!claset addIs [gcd_ind, nat_into_Ord RS Ord_0_lt]) 1);
qed "egcd_prop1";
(* if f divides a and b then f divides egcd(a,b) *)
goalw thy [dvd_def] "!!a. [| f dvd a; f dvd b; 0<b |] ==> f dvd (a mod b)";
-by (safe_tac (ZF_cs addSIs [mult_type, mod_type]));
+by (safe_tac (!claset addSIs [mult_type, mod_type]));
ren "m n" 1;
by (rtac (zero_lt_mult_iff RS iffD1 RS conjE) 1);
by (REPEAT_SOME assume_tac);
by (res_inst_tac
[("x", "(((m div n)#*n #+ m mod n) #- ((f#*m) div (f#*n)) #* n)")]
bexI 1);
-by (asm_simp_tac (arith_ss addsimps [diff_mult_distrib2, div_cancel,
+by (asm_simp_tac (!simpset addsimps [diff_mult_distrib2, div_cancel,
mult_mod_distrib, add_mult_distrib_left,
diff_add_inverse]) 1);
-by (asm_simp_tac arith_ss 1);
+by (Asm_simp_tac 1);
qed "dvd_mod";
@@ -147,29 +147,29 @@
by (rtac allI 1);
by (excluded_middle_tac "x=0" 1);
(* case x = 0 *)
-by (asm_simp_tac (arith_ss addsimps [egcd_0,dvd_refl,dvd_0_right,
+by (asm_simp_tac (!simpset addsimps [egcd_0,dvd_refl,dvd_0_right,
dvd_imp_nat2]) 2);
(* case x > 0 *)
-by (safe_tac ZF_cs);
-by (asm_simp_tac (arith_ss addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt,
+by (safe_tac (!claset));
+by (asm_simp_tac (!simpset addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt,
dvd_imp_nat2]) 1);
by (eres_inst_tac [("x","a mod x")] ballE 1);
by (asm_full_simp_tac
- (arith_ss addsimps [mod_less_divisor RS ltD, dvd_imp_nat2,
+ (!simpset addsimps [mod_less_divisor RS ltD, dvd_imp_nat2,
nat_into_Ord RS Ord_0_lt, egcd_lt_0]) 2);
-by (fast_tac (ZF_cs addSIs [dvd_mod, nat_into_Ord RS Ord_0_lt]) 1);
+by (fast_tac (!claset addSIs [dvd_mod, nat_into_Ord RS Ord_0_lt]) 1);
qed "egcd_prop2";
(* GCD PROOF : GCD exists and egcd fits the definition *)
goalw thy [gcd_def] "!!b. [| a: nat; b:nat |] ==> gcd(egcd(a,b), a, b)";
-by (asm_simp_tac (arith_ss addsimps [egcd_prop1]) 1);
-by (fast_tac (ZF_cs addIs [egcd_prop2 RS spec RS mp, dvd_imp_nat1]) 1);
+by (asm_simp_tac (!simpset addsimps [egcd_prop1]) 1);
+by (fast_tac (!claset addIs [egcd_prop2 RS spec RS mp, dvd_imp_nat1]) 1);
qed "gcd";
(* GCD is unique *)
goalw thy [gcd_def] "!!a. gcd(m,a,b) & gcd(n,a,b) --> m=n";
-by (fast_tac (ZF_cs addIs [dvd_anti_sym]) 1);
+by (fast_tac (!claset addIs [dvd_anti_sym]) 1);
qed "gcd_unique";
--- a/src/ZF/ex/Primrec.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Primrec.ML Fri Jan 03 15:01:55 1997 +0100
@@ -22,43 +22,38 @@
(** Useful special cases of evaluation ***)
-val pr_ss = arith_ss
- addsimps list.case_eqns
- addsimps [list_rec_Nil, list_rec_Cons,
- drop_0, drop_Nil, drop_succ_Cons,
- map_Nil, map_Cons]
- setsolver (type_auto_tac pr_typechecks);
+simpset := !simpset setsolver (type_auto_tac pr_typechecks);
goalw Primrec.thy [SC_def]
"!!x l. [| x:nat; l: list(nat) |] ==> SC ` (Cons(x,l)) = succ(x)";
-by (asm_simp_tac pr_ss 1);
+by (Asm_simp_tac 1);
qed "SC";
goalw Primrec.thy [CONST_def]
"!!l. [| l: list(nat) |] ==> CONST(k) ` l = k";
-by (asm_simp_tac pr_ss 1);
+by (Asm_simp_tac 1);
qed "CONST";
goalw Primrec.thy [PROJ_def]
"!!l. [| x: nat; l: list(nat) |] ==> PROJ(0) ` (Cons(x,l)) = x";
-by (asm_simp_tac pr_ss 1);
+by (Asm_simp_tac 1);
qed "PROJ_0";
goalw Primrec.thy [COMP_def]
"!!l. [| l: list(nat) |] ==> COMP(g,[f]) ` l = g` [f`l]";
-by (asm_simp_tac pr_ss 1);
+by (Asm_simp_tac 1);
qed "COMP_1";
goalw Primrec.thy [PREC_def]
"!!l. l: list(nat) ==> PREC(f,g) ` (Cons(0,l)) = f`l";
-by (asm_simp_tac pr_ss 1);
+by (Asm_simp_tac 1);
qed "PREC_0";
goalw Primrec.thy [PREC_def]
"!!l. [| x:nat; l: list(nat) |] ==> \
\ PREC(f,g) ` (Cons(succ(x),l)) = \
\ g ` Cons(PREC(f,g)`(Cons(x,l)), Cons(x,l))";
-by (asm_simp_tac pr_ss 1);
+by (Asm_simp_tac 1);
qed "PREC_succ";
(*** Inductive definition of the PR functions ***)
@@ -66,13 +61,12 @@
(* c: primrec ==> c: list(nat) -> nat *)
val primrec_into_fun = primrec.dom_subset RS subsetD;
-val pr_ss = pr_ss
- setsolver (type_auto_tac ([primrec_into_fun] @
- pr_typechecks @ primrec.intrs));
+simpset := !simpset setsolver (type_auto_tac ([primrec_into_fun] @
+ pr_typechecks @ primrec.intrs));
goalw Primrec.thy [ACK_def] "!!i. i:nat ==> ACK(i): primrec";
by (etac nat_induct 1);
-by (ALLGOALS (asm_simp_tac pr_ss));
+by (ALLGOALS Asm_simp_tac);
qed "ACK_in_primrec";
val ack_typechecks =
@@ -93,12 +87,12 @@
(*PROPERTY A 1*)
goalw Primrec.thy [ACK_def] "!!j. j:nat ==> ack(0,j) = succ(j)";
-by (asm_simp_tac (pr_ss addsimps [SC]) 1);
+by (asm_simp_tac (!simpset addsimps [SC]) 1);
qed "ack_0";
(*PROPERTY A 2*)
goalw Primrec.thy [ACK_def] "ack(succ(i), 0) = ack(i,1)";
-by (asm_simp_tac (pr_ss addsimps [CONST,PREC_0]) 1);
+by (asm_simp_tac (!simpset addsimps [CONST,PREC_0]) 1);
qed "ack_succ_0";
(*PROPERTY A 3*)
@@ -107,30 +101,28 @@
goalw Primrec.thy [ACK_def]
"!!i j. [| i:nat; j:nat |] ==> \
\ ack(succ(i), succ(j)) = ack(i, ack(succ(i), j))";
-by (asm_simp_tac (pr_ss addsimps [CONST,PREC_succ,COMP_1,PROJ_0]) 1);
+by (asm_simp_tac (!simpset addsimps [CONST,PREC_succ,COMP_1,PROJ_0]) 1);
qed "ack_succ_succ";
-val ack_ss =
- pr_ss addsimps [ack_0, ack_succ_0, ack_succ_succ,
- ack_type, nat_into_Ord];
+Addsimps [ack_0, ack_succ_0, ack_succ_succ, ack_type, nat_into_Ord];
(*PROPERTY A 4*)
goal Primrec.thy "!!i. i:nat ==> ALL j:nat. j < ack(i,j)";
by (etac nat_induct 1);
-by (asm_simp_tac ack_ss 1);
+by (Asm_simp_tac 1);
by (rtac ballI 1);
by (eres_inst_tac [("n","j")] nat_induct 1);
by (DO_GOAL [rtac (nat_0I RS nat_0_le RS lt_trans),
- asm_simp_tac ack_ss] 1);
+ Asm_simp_tac] 1);
by (DO_GOAL [etac (succ_leI RS lt_trans1),
- asm_simp_tac ack_ss] 1);
+ Asm_simp_tac] 1);
qed "lt_ack2_lemma";
bind_thm ("lt_ack2", (lt_ack2_lemma RS bspec));
(*PROPERTY A 5-, the single-step lemma*)
goal Primrec.thy "!!i j. [| i:nat; j:nat |] ==> ack(i,j) < ack(i, succ(j))";
by (etac nat_induct 1);
-by (ALLGOALS (asm_simp_tac (ack_ss addsimps [lt_ack2])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [lt_ack2])));
qed "ack_lt_ack_succ2";
(*PROPERTY A 5, monotonicity for < *)
@@ -153,7 +145,7 @@
goal Primrec.thy
"!!i j. [| i:nat; j:nat |] ==> ack(i, succ(j)) le ack(succ(i), j)";
by (nat_ind_tac "j" [] 1);
-by (ALLGOALS (asm_simp_tac ack_ss));
+by (ALLGOALS Asm_simp_tac);
by (rtac ack_le_mono2 1);
by (rtac (lt_ack2 RS succ_leI RS le_trans) 1);
by (REPEAT (ares_tac (ack_typechecks) 1));
@@ -185,13 +177,13 @@
(*PROPERTY A 8*)
goal Primrec.thy "!!j. j:nat ==> ack(1,j) = succ(succ(j))";
by (etac nat_induct 1);
-by (ALLGOALS (asm_simp_tac ack_ss));
+by (ALLGOALS Asm_simp_tac);
qed "ack_1";
(*PROPERTY A 9*)
goal Primrec.thy "!!j. j:nat ==> ack(succ(1),j) = succ(succ(succ(j#+j)))";
by (etac nat_induct 1);
-by (ALLGOALS (asm_simp_tac (ack_ss addsimps [ack_1, add_succ_right])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [ack_1, add_succ_right])));
qed "ack_2";
(*PROPERTY A 10*)
@@ -199,7 +191,7 @@
"!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \
\ ack(i1, ack(i2,j)) < ack(succ(succ(i1#+i2)), j)";
by (rtac (ack2_le_ack1 RSN (2,lt_trans2)) 1);
-by (asm_simp_tac ack_ss 1);
+by (Asm_simp_tac 1);
by (rtac (add_le_self RS ack_le_mono1 RS lt_trans1) 1);
by (rtac (add_le_self2 RS ack_lt_mono1 RS ack_lt_mono2) 5);
by (tc_tac []);
@@ -210,9 +202,9 @@
"!!i1 i2 j. [| i1:nat; i2:nat; j:nat |] ==> \
\ ack(i1,j) #+ ack(i2,j) < ack(succ(succ(succ(succ(i1#+i2)))), j)";
by (res_inst_tac [("j", "ack(succ(1), ack(i1 #+ i2, j))")] lt_trans 1);
-by (asm_simp_tac (ack_ss addsimps [ack_2]) 1);
+by (asm_simp_tac (!simpset addsimps [ack_2]) 1);
by (rtac (ack_nest_bound RS lt_trans2) 2);
-by (asm_simp_tac ack_ss 5);
+by (Asm_simp_tac 5);
by (rtac (add_le_mono RS leI RS leI) 1);
by (REPEAT (ares_tac ([add_le_self, add_le_self2, ack_le_mono1] @
ack_typechecks) 1));
@@ -225,45 +217,44 @@
\ i#+j < ack(succ(succ(succ(succ(k)))), j)";
by (res_inst_tac [("j", "ack(k,j) #+ ack(0,j)")] lt_trans 1);
by (rtac (ack_add_bound RS lt_trans2) 2);
-by (asm_simp_tac (ack_ss addsimps [add_0_right]) 5);
+by (asm_simp_tac (!simpset addsimps [add_0_right]) 5);
by (REPEAT (ares_tac ([add_lt_mono, lt_ack2] @ ack_typechecks) 1));
qed "ack_add_bound2";
(*** MAIN RESULT ***)
-val ack2_ss =
- ack_ss addsimps [list_add_Nil, list_add_Cons, list_add_type, nat_into_Ord];
+Addsimps [list_add_type, nat_into_Ord];
goalw Primrec.thy [SC_def]
"!!l. l: list(nat) ==> SC ` l < ack(1, list_add(l))";
by (etac list.elim 1);
-by (asm_simp_tac (ack2_ss addsimps [succ_iff]) 1);
-by (asm_simp_tac (ack2_ss addsimps [ack_1, add_le_self]) 1);
+by (asm_simp_tac (!simpset addsimps [succ_iff]) 1);
+by (asm_simp_tac (!simpset addsimps [ack_1, add_le_self]) 1);
qed "SC_case";
(*PROPERTY A 4'? Extra lemma needed for CONST case, constant functions*)
goal Primrec.thy "!!j. [| i:nat; j:nat |] ==> i < ack(i,j)";
by (etac nat_induct 1);
-by (asm_simp_tac (ack_ss addsimps [nat_0_le]) 1);
+by (asm_simp_tac (!simpset addsimps [nat_0_le]) 1);
by (etac ([succ_leI, ack_lt_ack_succ1] MRS lt_trans1) 1);
by (tc_tac []);
qed "lt_ack1";
goalw Primrec.thy [CONST_def]
"!!l. [| l: list(nat); k: nat |] ==> CONST(k) ` l < ack(k, list_add(l))";
-by (asm_simp_tac (ack2_ss addsimps [lt_ack1]) 1);
+by (asm_simp_tac (!simpset addsimps [lt_ack1]) 1);
qed "CONST_case";
goalw Primrec.thy [PROJ_def]
"!!l. l: list(nat) ==> ALL i:nat. PROJ(i) ` l < ack(0, list_add(l))";
-by (asm_simp_tac ack2_ss 1);
+by (Asm_simp_tac 1);
by (etac list.induct 1);
-by (asm_simp_tac (ack2_ss addsimps [nat_0_le]) 1);
-by (asm_simp_tac ack2_ss 1);
+by (asm_simp_tac (!simpset addsimps [nat_0_le]) 1);
+by (Asm_simp_tac 1);
by (rtac ballI 1);
by (eres_inst_tac [("n","x")] natE 1);
-by (asm_simp_tac (ack2_ss addsimps [add_le_self]) 1);
-by (asm_simp_tac ack2_ss 1);
+by (asm_simp_tac (!simpset addsimps [add_le_self]) 1);
+by (Asm_simp_tac 1);
by (etac (bspec RS lt_trans2) 1);
by (rtac (add_le_self2 RS succ_leI) 2);
by (tc_tac []);
@@ -280,10 +271,10 @@
\ list_add(map(%f. f ` l, fs)) < ack(k, list_add(l))";
by (etac list.induct 1);
by (DO_GOAL [res_inst_tac [("x","0")] bexI,
- asm_simp_tac (ack2_ss addsimps [lt_ack1, nat_0_le]),
+ asm_simp_tac (!simpset addsimps [lt_ack1, nat_0_le]),
resolve_tac nat_typechecks] 1);
-by (safe_tac ZF_cs);
-by (asm_simp_tac ack2_ss 1);
+by (safe_tac (!claset));
+by (Asm_simp_tac 1);
by (rtac (ballI RS bexI) 1);
by (rtac (add_lt_mono RS lt_trans) 1);
by (REPEAT (FIRSTGOAL (etac bspec)));
@@ -298,7 +289,7 @@
\ EX kf:nat. ALL l:list(nat). \
\ f`l < ack(kf, list_add(l))}) \
\ |] ==> EX k:nat. ALL l: list(nat). COMP(g,fs)`l < ack(k, list_add(l))";
-by (asm_simp_tac ZF_ss 1);
+by (Asm_simp_tac 1);
by (forward_tac [list_CollectD] 1);
by (etac (COMP_map_lemma RS bexE) 1);
by (rtac (ballI RS bexI) 1);
@@ -319,24 +310,24 @@
\ l: list(nat) \
\ |] ==> PREC(f,g)`l #+ list_add(l) < ack(succ(kf#+kg), list_add(l))";
by (etac list.elim 1);
-by (asm_simp_tac (ack2_ss addsimps [[nat_le_refl, lt_ack2] MRS lt_trans]) 1);
-by (asm_simp_tac ack2_ss 1);
+by (asm_simp_tac (!simpset addsimps [[nat_le_refl, lt_ack2] MRS lt_trans]) 1);
+by (Asm_simp_tac 1);
by (etac ssubst 1); (*get rid of the needless assumption*)
by (eres_inst_tac [("n","a")] nat_induct 1);
(*base case*)
-by (DO_GOAL [asm_simp_tac ack2_ss, rtac lt_trans, etac bspec,
+by (DO_GOAL [Asm_simp_tac, rtac lt_trans, etac bspec,
assume_tac, rtac (add_le_self RS ack_lt_mono1),
REPEAT o ares_tac (ack_typechecks)] 1);
(*ind step*)
-by (asm_simp_tac (ack2_ss addsimps [add_succ_right]) 1);
+by (asm_simp_tac (!simpset addsimps [add_succ_right]) 1);
by (rtac (succ_leI RS lt_trans1) 1);
by (res_inst_tac [("j", "g ` ?ll #+ ?mm")] lt_trans1 1);
by (etac bspec 2);
by (rtac (nat_le_refl RS add_le_mono) 1);
by (tc_tac []);
-by (asm_simp_tac (ack2_ss addsimps [add_le_self2]) 1);
+by (asm_simp_tac (!simpset addsimps [add_le_self2]) 1);
(*final part of the simplification*)
-by (asm_simp_tac ack2_ss 1);
+by (Asm_simp_tac 1);
by (rtac (add_le_self2 RS ack_le_mono1 RS lt_trans1) 1);
by (etac ack_lt_mono2 5);
by (tc_tac []);
@@ -361,7 +352,7 @@
goal Primrec.thy
"!!f. f:primrec ==> EX k:nat. ALL l:list(nat). f`l < ack(k, list_add(l))";
by (etac primrec.induct 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (DEPTH_SOLVE
(ares_tac ([SC_case, CONST_case, PROJ_case, COMP_case, PREC_case,
bexI, ballI] @ nat_typechecks) 1));
@@ -373,7 +364,7 @@
by (etac (ack_bounds_primrec RS bexE) 1);
by (rtac lt_irrefl 1);
by (dres_inst_tac [("x", "[x]")] bspec 1);
-by (asm_simp_tac ack2_ss 1);
-by (asm_full_simp_tac (ack2_ss addsimps [add_0_right]) 1);
+by (Asm_simp_tac 1);
+by (asm_full_simp_tac (!simpset addsimps [add_0_right]) 1);
qed "ack_not_primrec";
--- a/src/ZF/ex/PropLog.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/PropLog.ML Fri Jan 03 15:01:55 1997 +0100
@@ -18,13 +18,13 @@
goal PropLog.thy "prop_rec(Fls,b,c,d) = b";
by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac prop.con_defs);
-by (simp_tac rank_ss 1);
+by (Simp_tac 1);
qed "prop_rec_Fls";
goal PropLog.thy "prop_rec(#v,b,c,d) = c(v)";
by (rtac (prop_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac prop.con_defs);
-by (simp_tac rank_ss 1);
+by (Simp_tac 1);
qed "prop_rec_Var";
goal PropLog.thy "prop_rec(p=>q,b,c,d) = \
@@ -34,45 +34,43 @@
by (simp_tac rank_ss 1);
qed "prop_rec_Imp";
-val prop_rec_ss =
- arith_ss addsimps [prop_rec_Fls, prop_rec_Var, prop_rec_Imp];
+Addsimps [prop_rec_Fls, prop_rec_Var, prop_rec_Imp];
(*** Semantics of propositional logic ***)
(** The function is_true **)
goalw PropLog.thy [is_true_def] "is_true(Fls,t) <-> False";
-by (simp_tac (prop_rec_ss addsimps [one_not_0 RS not_sym]) 1);
+by (simp_tac (!simpset addsimps [one_not_0 RS not_sym]) 1);
qed "is_true_Fls";
goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t";
-by (simp_tac (prop_rec_ss addsimps [one_not_0 RS not_sym]
+by (simp_tac (!simpset addsimps [one_not_0 RS not_sym]
setloop (split_tac [expand_if])) 1);
qed "is_true_Var";
goalw PropLog.thy [is_true_def]
"is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))";
-by (simp_tac (prop_rec_ss setloop (split_tac [expand_if])) 1);
+by (simp_tac (!simpset setloop (split_tac [expand_if])) 1);
qed "is_true_Imp";
(** The function hyps **)
goalw PropLog.thy [hyps_def] "hyps(Fls,t) = 0";
-by (simp_tac prop_rec_ss 1);
+by (Simp_tac 1);
qed "hyps_Fls";
goalw PropLog.thy [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}";
-by (simp_tac prop_rec_ss 1);
+by (Simp_tac 1);
qed "hyps_Var";
goalw PropLog.thy [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)";
-by (simp_tac prop_rec_ss 1);
+by (Simp_tac 1);
qed "hyps_Imp";
-val prop_ss = prop_rec_ss
- addsimps prop.intrs
- addsimps [is_true_Fls, is_true_Var, is_true_Imp,
- hyps_Fls, hyps_Var, hyps_Imp];
+Addsimps prop.intrs;
+Addsimps [is_true_Fls, is_true_Var, is_true_Imp,
+ hyps_Fls, hyps_Var, hyps_Imp];
(*** Proof theory of propositional logic ***)
@@ -119,11 +117,11 @@
(*The deduction theorem*)
goal PropLog.thy "!!p q H. [| cons(p,H) |- q; p:prop |] ==> H |- p=>q";
by (etac thms.induct 1);
-by (fast_tac (ZF_cs addIs [thms_I, thms.H RS weaken_right]) 1);
-by (fast_tac (ZF_cs addIs [thms.K RS weaken_right]) 1);
-by (fast_tac (ZF_cs addIs [thms.S RS weaken_right]) 1);
-by (fast_tac (ZF_cs addIs [thms.DN RS weaken_right]) 1);
-by (fast_tac (ZF_cs addIs [thms.S RS thms_MP RS thms_MP]) 1);
+by (fast_tac (!claset addIs [thms_I, thms.H RS weaken_right]) 1);
+by (fast_tac (!claset addIs [thms.K RS weaken_right]) 1);
+by (fast_tac (!claset addIs [thms.S RS weaken_right]) 1);
+by (fast_tac (!claset addIs [thms.DN RS weaken_right]) 1);
+by (fast_tac (!claset addIs [thms.S RS thms_MP RS thms_MP]) 1);
qed "deduction";
@@ -145,8 +143,8 @@
(*Soundness of the rules wrt truth-table semantics*)
goalw PropLog.thy [logcon_def] "!!H. H |- p ==> H |= p";
by (etac thms.induct 1);
-by (fast_tac (ZF_cs addSDs [is_true_Imp RS iffD1 RS mp]) 5);
-by (ALLGOALS (asm_simp_tac prop_ss));
+by (fast_tac (!claset addSDs [is_true_Imp RS iffD1 RS mp]) 5);
+by (ALLGOALS Asm_simp_tac);
qed "soundness";
(*** Towards the completeness proof ***)
@@ -175,10 +173,10 @@
"p: prop ==> hyps(p,t) |- if(is_true(p,t), p, p=>Fls)";
by (rtac (expand_if RS iffD2) 1);
by (rtac (major RS prop.induct) 1);
-by (ALLGOALS (asm_simp_tac (prop_ss addsimps [thms_I, thms.H])));
-by (safe_tac (ZF_cs addSEs [Fls_Imp RS weaken_left_Un1,
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [thms_I, thms.H])));
+by (safe_tac (!claset addSEs [Fls_Imp RS weaken_left_Un1,
Fls_Imp RS weaken_left_Un2]));
-by (ALLGOALS (fast_tac (ZF_cs addIs [weaken_left_Un1, weaken_left_Un2,
+by (ALLGOALS (fast_tac (!claset addIs [weaken_left_Un1, weaken_left_Un2,
weaken_right, Imp_Fls])));
qed "hyps_thms_if";
@@ -187,7 +185,7 @@
"[| p: prop; 0 |= p |] ==> hyps(p,t) |- p";
by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN
rtac (premp RS hyps_thms_if) 2);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "logcon_thms_p";
(*For proving certain theorems in our new propositional logic*)
@@ -217,11 +215,11 @@
val [major] = goal PropLog.thy
"p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})";
by (rtac (major RS prop.induct) 1);
-by (simp_tac prop_ss 1);
-by (asm_simp_tac (prop_ss setloop (split_tac [expand_if])) 1);
-by (fast_tac (ZF_cs addSEs prop.free_SEs) 1);
-by (asm_simp_tac prop_ss 1);
-by (fast_tac ZF_cs 1);
+by (Simp_tac 1);
+by (asm_simp_tac (!simpset setloop (split_tac [expand_if])) 1);
+by (fast_tac (!claset addSEs prop.free_SEs) 1);
+by (Asm_simp_tac 1);
+by (Fast_tac 1);
qed "hyps_Diff";
(*For the case hyps(p,t)-cons(#v => Fls,Y) |- p;
@@ -229,21 +227,21 @@
val [major] = goal PropLog.thy
"p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})";
by (rtac (major RS prop.induct) 1);
-by (simp_tac prop_ss 1);
-by (asm_simp_tac (prop_ss setloop (split_tac [expand_if])) 1);
-by (fast_tac (ZF_cs addSEs prop.free_SEs) 1);
-by (asm_simp_tac prop_ss 1);
-by (fast_tac ZF_cs 1);
+by (Simp_tac 1);
+by (asm_simp_tac (!simpset setloop (split_tac [expand_if])) 1);
+by (fast_tac (!claset addSEs prop.free_SEs) 1);
+by (Asm_simp_tac 1);
+by (Fast_tac 1);
qed "hyps_cons";
(** Two lemmas for use with weaken_left **)
-goal ZF.thy "B-C <= cons(a, B-cons(a,C))";
-by (fast_tac ZF_cs 1);
+goal upair.thy "B-C <= cons(a, B-cons(a,C))";
+by (Fast_tac 1);
qed "cons_Diff_same";
-goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
-by (fast_tac ZF_cs 1);
+goal upair.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
+by (Fast_tac 1);
qed "cons_Diff_subset2";
(*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls;
@@ -251,9 +249,10 @@
val [major] = goal PropLog.thy
"p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})";
by (rtac (major RS prop.induct) 1);
-by (asm_simp_tac (prop_ss addsimps (Fin.intrs @ [UN_I, cons_iff])
+by (asm_simp_tac (!simpset addsimps [UN_I]
setloop (split_tac [expand_if])) 2);
-by (ALLGOALS (asm_simp_tac (prop_ss addsimps [Un_0, Fin.emptyI, Fin_UnI])));
+by (ALLGOALS Asm_simp_tac);
+by (fast_tac (!claset addIs Fin.intrs) 1);
qed "hyps_finite";
val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left;
@@ -263,8 +262,8 @@
val [premp,sat] = goal PropLog.thy
"[| p: prop; 0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p";
by (rtac (premp RS hyps_finite RS Fin_induct) 1);
-by (simp_tac (prop_ss addsimps [premp, sat, logcon_thms_p, Diff_0]) 1);
-by (safe_tac ZF_cs);
+by (simp_tac (!simpset addsimps [premp, sat, logcon_thms_p, Diff_0]) 1);
+by (safe_tac (!claset));
(*Case hyps(p,t)-cons(#v,Y) |- p *)
by (rtac thms_excluded_middle_rule 1);
by (etac prop.Var_I 3);
@@ -291,22 +290,22 @@
(*A semantic analogue of the Deduction Theorem*)
goalw PropLog.thy [logcon_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q";
-by (simp_tac prop_ss 1);
-by (fast_tac ZF_cs 1);
+by (Simp_tac 1);
+by (Fast_tac 1);
qed "logcon_Imp";
goal PropLog.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p";
by (etac Fin_induct 1);
-by (safe_tac (ZF_cs addSIs [completeness_0]));
+by (safe_tac (!claset addSIs [completeness_0]));
by (rtac (weaken_left_cons RS thms_MP) 1);
-by (fast_tac (ZF_cs addSIs (logcon_Imp::prop.intrs)) 1);
+by (fast_tac (!claset addSIs (logcon_Imp::prop.intrs)) 1);
by (fast_tac thms_cs 1);
qed "completeness_lemma";
val completeness = completeness_lemma RS bspec RS mp;
val [finite] = goal PropLog.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop";
-by (fast_tac (ZF_cs addSEs [soundness, finite RS completeness,
+by (fast_tac (!claset addSEs [soundness, finite RS completeness,
thms_in_pl]) 1);
qed "thms_iff";
--- a/src/ZF/ex/Ramsey.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Ramsey.ML Fri Jan 03 15:01:55 1997 +0100
@@ -23,42 +23,41 @@
(*** Cliques and Independent sets ***)
goalw Ramsey.thy [Clique_def] "Clique(0,V,E)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Clique0";
goalw Ramsey.thy [Clique_def]
"!!C V E. [| Clique(C,V',E); V'<=V |] ==> Clique(C,V,E)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Clique_superset";
goalw Ramsey.thy [Indept_def] "Indept(0,V,E)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Indept0";
val prems = goalw Ramsey.thy [Indept_def]
"!!I V E. [| Indept(I,V',E); V'<=V |] ==> Indept(I,V,E)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Indept_superset";
(*** Atleast ***)
goalw Ramsey.thy [Atleast_def, inj_def, Pi_def, function_def] "Atleast(0,A)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Atleast0";
goalw Ramsey.thy [Atleast_def]
"!!m A. Atleast(succ(m),A) ==> EX x:A. Atleast(m, A-{x})";
-by (fast_tac (ZF_cs addEs [inj_is_fun RS apply_type, inj_succ_restrict]) 1);
+by (fast_tac (!claset addEs [inj_is_fun RS apply_type, inj_succ_restrict]) 1);
qed "Atleast_succD";
goalw Ramsey.thy [Atleast_def]
"!!n A. [| Atleast(n,A); A<=B |] ==> Atleast(n,B)";
-by (fast_tac (ZF_cs addEs [inj_weaken_type]) 1);
+by (fast_tac (!claset addEs [inj_weaken_type]) 1);
qed "Atleast_superset";
-val prems = goalw Ramsey.thy [Atleast_def,succ_def]
- "[| Atleast(m,B); b~: B |] ==> Atleast(succ(m), cons(b,B))";
-by (cut_facts_tac prems 1);
+goalw Ramsey.thy [Atleast_def,succ_def]
+ "!!m. [| Atleast(m,B); b~: B |] ==> Atleast(succ(m), cons(b,B))";
by (etac exE 1);
by (rtac exI 1);
by (etac inj_extend 1);
@@ -66,12 +65,11 @@
by (assume_tac 1);
qed "Atleast_succI";
-val prems = goal Ramsey.thy
- "[| Atleast(m, B-{x}); x: B |] ==> Atleast(succ(m), B)";
-by (cut_facts_tac prems 1);
+goal Ramsey.thy
+ "!!m. [| Atleast(m, B-{x}); x: B |] ==> Atleast(succ(m), B)";
by (etac (Atleast_succI RS Atleast_superset) 1);
-by (fast_tac ZF_cs 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
+by (Fast_tac 1);
qed "Atleast_Diff_succI";
(*** Main Cardinality Lemma ***)
@@ -83,14 +81,14 @@
\ ALL n: nat. ALL A B. Atleast((m#+n) #- succ(0), A Un B) --> \
\ Atleast(m,A) | Atleast(n,B)";
by (nat_ind_tac "m" prems 1);
-by (fast_tac (ZF_cs addSIs [Atleast0]) 1);
-by (asm_simp_tac arith_ss 1);
+by (fast_tac (!claset addSIs [Atleast0]) 1);
+by (Asm_simp_tac 1);
by (rtac ballI 1);
by (rename_tac "n" 1); (*simplifier does NOT preserve bound names!*)
by (nat_ind_tac "n" [] 1);
-by (fast_tac (ZF_cs addSIs [Atleast0]) 1);
-by (asm_simp_tac (arith_ss addsimps [add_succ_right]) 1);
-by (safe_tac ZF_cs);
+by (fast_tac (!claset addSIs [Atleast0]) 1);
+by (asm_simp_tac (!simpset addsimps [add_succ_right]) 1);
+by (safe_tac (!claset));
by (etac (Atleast_succD RS bexE) 1);
by (etac UnE 1);
(**case x:B. Instantiate the 'ALL A B' induction hypothesis. **)
@@ -99,7 +97,7 @@
(*cases Atleast(succ(m1),A) and Atleast(succ(n1),B)*)
by (REPEAT (eresolve_tac [asm_rl, notE, Atleast_Diff_succI] 3));
(*proving the condition*)
-by (etac Atleast_superset 2 THEN fast_tac ZF_cs 2);
+by (etac Atleast_superset 2 THEN Fast_tac 2);
(**case x:A. Instantiate the 'ALL n:nat. ALL A B' induction hypothesis. **)
by (dres_inst_tac [("x2","succ(n1)"), ("x1","A-{x}"), ("x","B")]
(bspec RS spec RS spec) 1);
@@ -108,8 +106,8 @@
(*cases Atleast(succ(m1),A) and Atleast(succ(n1),B)*)
by (REPEAT (eresolve_tac [asm_rl, Atleast_Diff_succI, notE] 2));
(*proving the condition*)
-by (asm_simp_tac (arith_ss addsimps [add_succ_right]) 1);
-by (etac Atleast_superset 1 THEN fast_tac ZF_cs 1);
+by (asm_simp_tac (!simpset addsimps [add_succ_right]) 1);
+by (etac Atleast_superset 1 THEN Fast_tac 1);
qed "pigeon2_lemma";
(* [| m:nat; n:nat; Atleast(m #+ n #- succ(0), A Un B) |] ==>
@@ -122,11 +120,11 @@
(** Base cases of induction; they now admit ANY Ramsey number **)
goalw Ramsey.thy [Ramsey_def] "Ramsey(n,0,j)";
-by (fast_tac (ZF_cs addIs [Clique0,Atleast0]) 1);
+by (fast_tac (!claset addIs [Clique0,Atleast0]) 1);
qed "Ramsey0j";
goalw Ramsey.thy [Ramsey_def] "Ramsey(n,i,0)";
-by (fast_tac (ZF_cs addIs [Indept0,Atleast0]) 1);
+by (fast_tac (!claset addIs [Indept0,Atleast0]) 1);
qed "Ramseyi0";
(** Lemmas for induction step **)
@@ -137,10 +135,10 @@
"[| Atleast(m #+ n, A); m: nat; n: nat |] ==> \
\ Atleast(succ(m), {x:A. ~P(x)}) | Atleast(n, {x:A. P(x)})";
by (rtac (nat_succI RS pigeon2) 1);
-by (simp_tac (arith_ss addsimps prems) 3);
+by (simp_tac (!simpset addsimps prems) 3);
by (rtac Atleast_superset 3);
by (REPEAT (resolve_tac prems 1));
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Atleast_partition";
(*For the Atleast part, proves ~(a:I) from the second premise!*)
@@ -149,7 +147,7 @@
\ Atleast(j,I) |] ==> \
\ Indept(cons(a,I), V, E) & Atleast(succ(j), cons(a,I))";
by (cut_facts_tac prems 1);
-by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*34 secs*)
+by (fast_tac (!claset addSEs [Atleast_succI]) 1); (*34 secs*)
qed "Indept_succ";
val prems = goalw Ramsey.thy [Symmetric_def,Clique_def]
@@ -157,7 +155,7 @@
\ Atleast(j,C) |] ==> \
\ Clique(cons(a,C), V, E) & Atleast(succ(j), cons(a,C))";
by (cut_facts_tac prems 1);
-by (fast_tac (ZF_cs addSEs [Atleast_succI]) 1); (*41 secs*)
+by (fast_tac (!claset addSEs [Atleast_succI]) 1); (*41 secs*)
qed "Clique_succ";
(** Induction step **)
@@ -167,24 +165,24 @@
"[| Ramsey(succ(m), succ(i), j); Ramsey(n, i, succ(j)); \
\ m: nat; n: nat |] ==> \
\ Ramsey(succ(m#+n), succ(i), succ(j))";
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (etac (Atleast_succD RS bexE) 1);
by (eres_inst_tac [("P1","%z.<x,z>:E")] (Atleast_partition RS disjE) 1);
by (REPEAT (resolve_tac prems 1));
(*case m*)
by (rtac (ram1 RS spec RS spec RS mp RS disjE) 1);
-by (fast_tac ZF_cs 1);
-by (fast_tac (ZF_cs addEs [Clique_superset]) 1); (*easy -- given a Clique*)
-by (safe_tac ZF_cs);
+by (Fast_tac 1);
+by (fast_tac (!claset addEs [Clique_superset]) 1); (*easy -- given a Clique*)
+by (safe_tac (!claset));
by (eresolve_tac (swapify [exI]) 1); (*ignore main EX quantifier*)
by (REPEAT (ares_tac [Indept_succ] 1)); (*make a bigger Indept*)
(*case n*)
by (rtac (ram2 RS spec RS spec RS mp RS disjE) 1);
-by (fast_tac ZF_cs 1);
-by (safe_tac ZF_cs);
+by (Fast_tac 1);
+by (safe_tac (!claset));
by (rtac exI 1);
by (REPEAT (ares_tac [Clique_succ] 1)); (*make a bigger Clique*)
-by (fast_tac (ZF_cs addEs [Indept_superset]) 1); (*easy -- given an Indept*)
+by (fast_tac (!claset addEs [Indept_superset]) 1); (*easy -- given an Indept*)
qed "Ramsey_step_lemma";
@@ -194,16 +192,17 @@
val prems = goal Ramsey.thy
"i: nat ==> ALL j: nat. EX n:nat. Ramsey(succ(n), i, j)";
by (nat_ind_tac "i" prems 1);
-by (fast_tac (ZF_cs addSIs [nat_0I,Ramsey0j]) 1);
+by (fast_tac (!claset addSIs [Ramsey0j]) 1);
by (rtac ballI 1);
by (nat_ind_tac "j" [] 1);
-by (fast_tac (ZF_cs addSIs [nat_0I,Ramseyi0]) 1);
-by (fast_tac (ZF_cs addSIs [nat_succI,add_type,Ramsey_step_lemma]) 1);
+by (fast_tac (!claset addSIs [Ramseyi0]) 1);
+by (fast_tac (!claset addSDs [bspec]
+ addSIs [nat_succI,add_type,Ramsey_step_lemma]) 1);
qed "ramsey_lemma";
(*Final statement in a tidy form, without succ(...) *)
goal Ramsey.thy "!!i j. [| i: nat; j: nat |] ==> EX n:nat. Ramsey(n,i,j)";
-by (best_tac (ZF_cs addDs [ramsey_lemma] addSIs [nat_succI]) 1);
+by (best_tac (!claset addDs [ramsey_lemma] addSIs [nat_succI]) 1);
qed "ramsey";
(*Compute Ramsey numbers according to proof above -- which, actually,
--- a/src/ZF/ex/Rmap.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Rmap.ML Fri Jan 03 15:01:55 1997 +0100
@@ -18,8 +18,8 @@
val Nil_rmap_case = rmap.mk_cases list.con_defs "<Nil,zs> : rmap(r)"
and Cons_rmap_case = rmap.mk_cases list.con_defs "<Cons(x,xs),zs> : rmap(r)";
-val rmap_cs = ZF_cs addIs rmap.intrs
- addSEs [Nil_rmap_case, Cons_rmap_case];
+AddIs rmap.intrs;
+AddSEs [Nil_rmap_case, Cons_rmap_case];
goal Rmap.thy "!!r. r <= A*B ==> rmap(r) <= list(A)*list(B)";
by (rtac (rmap.dom_subset RS subset_trans) 1);
@@ -31,13 +31,13 @@
"!!r. A <= domain(r) ==> list(A) <= domain(rmap(r))";
by (rtac subsetI 1);
by (etac list.induct 1);
-by (ALLGOALS (fast_tac rmap_cs));
+by (ALLGOALS Fast_tac);
qed "rmap_total";
goalw Rmap.thy [function_def] "!!r. function(r) ==> function(rmap(r))";
by (rtac (impI RS allI RS allI) 1);
by (etac rmap.induct 1);
-by (ALLGOALS (fast_tac rmap_cs));
+by (ALLGOALS Fast_tac);
qed "rmap_functional";
@@ -45,11 +45,11 @@
goal Rmap.thy "!!f. f: A->B ==> rmap(f): list(A)->list(B)";
by (asm_full_simp_tac
- (ZF_ss addsimps [Pi_iff, rmap_rel_type, rmap_functional, rmap_total]) 1);
+ (!simpset addsimps [Pi_iff, rmap_rel_type, rmap_functional, rmap_total]) 1);
qed "rmap_fun_type";
goalw Rmap.thy [apply_def] "rmap(f)`Nil = Nil";
-by (fast_tac (rmap_cs addIs [the_equality]) 1);
+by (fast_tac (!claset addIs [the_equality]) 1);
qed "rmap_Nil";
goal Rmap.thy "!!f. [| f: A->B; x: A; xs: list(A) |] ==> \
--- a/src/ZF/ex/TF.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/TF.ML Fri Jan 03 15:01:55 1997 +0100
@@ -29,7 +29,7 @@
goal TF.thy "tree(A) Un forest(A) = tree_forest(A)";
by (safe_tac (subset_cs addSIs [equalityI, tree_subset_TF, forest_subset_TF]));
-by (fast_tac (ZF_cs addSIs tree_forest.intrs addEs [tree_forest.elim]) 1);
+by (fast_tac (!claset addSIs tree_forest.intrs addEs [tree_forest.elim]) 1);
qed "TF_equals_Un";
(** NOT useful, but interesting... **)
@@ -72,7 +72,7 @@
goal TF.thy "TF_rec(Fnil, b, c, d) = c";
by (rtac (TF_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac tree_forest.con_defs);
-by (simp_tac rank_ss 1);
+by (Simp_tac 1);
qed "TF_rec_Fnil";
goal TF.thy "TF_rec(Fcons(t,f), b, c, d) = \
@@ -82,9 +82,7 @@
by (simp_tac rank_ss 1);
qed "TF_rec_Fcons";
-(*list_ss includes list operations as well as arith_ss*)
-val TF_rec_ss = list_ss addsimps
- [TF_rec_Tcons, TF_rec_Fnil, TF_rec_Fcons, TconsI, FnilI, FconsI];
+Addsimps [TF_rec_Tcons, TF_rec_Fnil, TF_rec_Fcons, TconsI, FnilI, FconsI];
(** Type checking **)
@@ -97,7 +95,7 @@
\ |] ==> d(t,f,r1,r2): C(Fcons(t,f)) \
\ |] ==> TF_rec(z,b,c,d) : C(z)";
by (rtac (major RS tree_forest.induct) 1);
-by (ALLGOALS (asm_simp_tac (TF_rec_ss addsimps prems)));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "TF_rec_type";
(*Mutually recursive version*)
@@ -111,7 +109,7 @@
\ (ALL f: forest(A). TF_rec(f,b,c,d) : D(f))";
by (rewtac Ball_def);
by (rtac tree_forest.mutual_induct 1);
-by (ALLGOALS (asm_simp_tac (TF_rec_ss addsimps prems)));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps prems)));
qed "tree_forest_rec_type";
@@ -143,6 +141,7 @@
val [list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons] =
TF_recs list_of_TF_def;
+Addsimps [list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons];
goalw TF.thy [list_of_TF_def]
"!!z A. z: tree_forest(A) ==> list_of_TF(z) : list(tree(A))";
@@ -150,6 +149,7 @@
qed "list_of_TF_type";
val [TF_of_list_Nil,TF_of_list_Cons] = list_recs TF_of_list_def;
+Addsimps [TF_of_list_Nil,TF_of_list_Cons];
goalw TF.thy [TF_of_list_def]
"!!l A. l: list(tree(A)) ==> TF_of_list(l) : forest(A)";
@@ -160,6 +160,7 @@
(** TF_map **)
val [TF_map_Tcons, TF_map_Fnil, TF_map_Fcons] = TF_recs TF_map_def;
+Addsimps [TF_map_Tcons, TF_map_Fnil, TF_map_Fcons];
val prems = goalw TF.thy [TF_map_def]
"[| !!x. x: A ==> h(x): B |] ==> \
@@ -173,6 +174,7 @@
(** TF_size **)
val [TF_size_Tcons, TF_size_Fnil, TF_size_Fcons] = TF_recs TF_size_def;
+Addsimps [TF_size_Tcons, TF_size_Fnil, TF_size_Fcons];
goalw TF.thy [TF_size_def]
"!!z A. z: tree_forest(A) ==> TF_size(z) : nat";
@@ -184,6 +186,7 @@
val [TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons] =
TF_recs TF_preorder_def;
+Addsimps [TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons];
goalw TF.thy [TF_preorder_def]
"!!z A. z: tree_forest(A) ==> TF_preorder(z) : list(A)";
@@ -200,16 +203,7 @@
[TconsI, FnilI, FconsI, treeI, forestI,
list_of_TF_type, TF_map_type, TF_size_type, TF_preorder_type];
-val TF_rewrites =
- [TF_rec_Tcons, TF_rec_Fnil, TF_rec_Fcons,
- list_of_TF_Tcons, list_of_TF_Fnil, list_of_TF_Fcons,
- TF_of_list_Nil,TF_of_list_Cons,
- TF_map_Tcons, TF_map_Fnil, TF_map_Fcons,
- TF_size_Tcons, TF_size_Fnil, TF_size_Fcons,
- TF_preorder_Tcons, TF_preorder_Fnil, TF_preorder_Fcons];
-
-val TF_ss = list_ss addsimps TF_rewrites
- setsolver type_auto_tac (list_typechecks@TF_typechecks);
+simpset := !simpset setsolver type_auto_tac (list_typechecks@TF_typechecks);
(** theorems about list_of_TF and TF_of_list **)
@@ -225,26 +219,26 @@
goal TF.thy "!!f A. f: forest(A) ==> TF_of_list(list_of_TF(f)) = f";
by (etac forest_induct 1);
-by (ALLGOALS (asm_simp_tac TF_ss));
+by (ALLGOALS Asm_simp_tac);
qed "forest_iso";
goal TF.thy
"!!ts. ts: list(tree(A)) ==> list_of_TF(TF_of_list(ts)) = ts";
by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac TF_ss));
+by (ALLGOALS Asm_simp_tac);
qed "tree_list_iso";
(** theorems about TF_map **)
goal TF.thy "!!z A. z: tree_forest(A) ==> TF_map(%u.u, z) = z";
by (etac tree_forest.induct 1);
-by (ALLGOALS (asm_simp_tac TF_ss));
+by (ALLGOALS Asm_simp_tac);
qed "TF_map_ident";
goal TF.thy
"!!z A. z: tree_forest(A) ==> TF_map(h, TF_map(j,z)) = TF_map(%u.h(j(u)), z)";
by (etac tree_forest.induct 1);
-by (ALLGOALS (asm_simp_tac TF_ss));
+by (ALLGOALS Asm_simp_tac);
qed "TF_map_compose";
(** theorems about TF_size **)
@@ -252,13 +246,13 @@
goal TF.thy
"!!z A. z: tree_forest(A) ==> TF_size(TF_map(h,z)) = TF_size(z)";
by (etac tree_forest.induct 1);
-by (ALLGOALS (asm_simp_tac TF_ss));
+by (ALLGOALS Asm_simp_tac);
qed "TF_size_TF_map";
goal TF.thy
"!!z A. z: tree_forest(A) ==> TF_size(z) = length(TF_preorder(z))";
by (etac tree_forest.induct 1);
-by (ALLGOALS (asm_simp_tac (TF_ss addsimps [length_app])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [length_app])));
qed "TF_size_length";
(** theorems about TF_preorder **)
@@ -266,5 +260,5 @@
goal TF.thy "!!z A. z: tree_forest(A) ==> \
\ TF_preorder(TF_map(h,z)) = map(h, TF_preorder(z))";
by (etac tree_forest.induct 1);
-by (ALLGOALS (asm_simp_tac (TF_ss addsimps [map_app_distrib])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [map_app_distrib])));
qed "TF_preorder_TF_map";
--- a/src/ZF/ex/Term.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/Term.ML Fri Jan 03 15:01:55 1997 +0100
@@ -52,7 +52,7 @@
goalw Term.thy (term.defs@term.con_defs) "term(univ(A)) <= univ(A)";
by (rtac lfp_lowerbound 1);
by (rtac (A_subset_univ RS univ_mono) 2);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
by (REPEAT (ares_tac [Pair_in_univ, list_univ RS subsetD] 1));
qed "term_univ";
@@ -71,11 +71,11 @@
"[| l: list(A); Ord(i) |] ==> \
\ rank(l)<i --> map(%z. (lam x:Vset(i).h(x)) ` z, l) = map(h,l)";
by (rtac (major RS list.induct) 1);
-by (simp_tac list_ss 1);
+by (Simp_tac 1);
by (rtac impI 1);
by (forward_tac [rank_Cons1 RS lt_trans] 1);
by (dtac (rank_Cons2 RS lt_trans) 1);
-by (asm_simp_tac (list_ss addsimps [ordi, VsetI]) 1);
+by (asm_simp_tac (!simpset addsimps [ordi, VsetI]) 1);
qed "map_lemma";
(*Typing premise is necessary to invoke map_lemma*)
@@ -84,8 +84,7 @@
\ term_rec(Apply(a,ts), d) = d(a, ts, map (%z. term_rec(z,d), ts))";
by (rtac (term_rec_def RS def_Vrec RS trans) 1);
by (rewrite_goals_tac term.con_defs);
-val term_rec_ss = ZF_ss addsimps [Ord_rank, rank_pair2, prem RS map_lemma];
-by (simp_tac term_rec_ss 1);
+by (simp_tac (!simpset addsimps [Ord_rank, rank_pair2, prem RS map_lemma]) 1);;
qed "term_rec";
(*Slightly odd typing condition on r in the second premise!*)
@@ -100,7 +99,7 @@
by (stac term_rec 1);
by (REPEAT (ares_tac prems 1));
by (etac list.induct 1);
-by (ALLGOALS (asm_simp_tac (list_ss addsimps [term_rec])));
+by (ALLGOALS (asm_simp_tac (!simpset addsimps [term_rec])));
by (etac CollectE 1);
by (REPEAT (ares_tac [list.Cons_I, UN_I] 1));
qed "term_rec_type";
@@ -172,7 +171,7 @@
reflect_type, preorder_type];
(*map_type2 and term_map_type2 instantiate variables*)
-val term_ss = list_ss
+simpset := !simpset
addsimps [term_rec, term_map, term_size, reflect, preorder]
setsolver type_auto_tac (list_typechecks@term_typechecks);
@@ -181,19 +180,19 @@
goal Term.thy "!!t A. t: term(A) ==> term_map(%u.u, t) = t";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [map_ident]) 1);
+by (asm_simp_tac (!simpset addsimps [map_ident]) 1);
qed "term_map_ident";
goal Term.thy
"!!t A. t: term(A) ==> term_map(f, term_map(g,t)) = term_map(%u.f(g(u)), t)";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [map_compose]) 1);
+by (asm_simp_tac (!simpset addsimps [map_compose]) 1);
qed "term_map_compose";
goal Term.thy
"!!t A. t: term(A) ==> term_map(f, reflect(t)) = reflect(term_map(f,t))";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [rev_map_distrib RS sym, map_compose]) 1);
+by (asm_simp_tac (!simpset addsimps [rev_map_distrib RS sym, map_compose]) 1);
qed "term_map_reflect";
@@ -202,18 +201,18 @@
goal Term.thy
"!!t A. t: term(A) ==> term_size(term_map(f,t)) = term_size(t)";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [map_compose]) 1);
+by (asm_simp_tac (!simpset addsimps [map_compose]) 1);
qed "term_size_term_map";
goal Term.thy "!!t A. t: term(A) ==> term_size(reflect(t)) = term_size(t)";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [rev_map_distrib RS sym, map_compose,
+by (asm_simp_tac (!simpset addsimps [rev_map_distrib RS sym, map_compose,
list_add_rev]) 1);
qed "term_size_reflect";
goal Term.thy "!!t A. t: term(A) ==> term_size(t) = length(preorder(t))";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [length_flat, map_compose]) 1);
+by (asm_simp_tac (!simpset addsimps [length_flat, map_compose]) 1);
qed "term_size_length";
@@ -221,7 +220,7 @@
goal Term.thy "!!t A. t: term(A) ==> reflect(reflect(t)) = t";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [rev_map_distrib, map_compose,
+by (asm_simp_tac (!simpset addsimps [rev_map_distrib, map_compose,
map_ident, rev_rev_ident]) 1);
qed "reflect_reflect_ident";
@@ -231,7 +230,7 @@
goal Term.thy
"!!t A. t: term(A) ==> preorder(term_map(f,t)) = map(f, preorder(t))";
by (etac term_induct_eqn 1);
-by (asm_simp_tac (term_ss addsimps [map_compose, map_flat]) 1);
+by (asm_simp_tac (!simpset addsimps [map_compose, map_flat]) 1);
qed "preorder_term_map";
(** preorder(reflect(t)) = rev(postorder(t)) **)
--- a/src/ZF/ex/misc.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/ex/misc.ML Fri Jan 03 15:01:55 1997 +0100
@@ -9,13 +9,14 @@
writeln"ZF/ex/misc";
+set_current_thy"Perm";
(*Example 12 (credited to Peter Andrews) from
W. Bledsoe. A Maximal Method for Set Variables in Automatic Theorem-proving.
In: J. Hayes and D. Michie and L. Mikulich, eds. Machine Intelligence 9.
Ellis Horwood, 53-100 (1979). *)
-goal ZF.thy "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
-by (best_tac ZF_cs 1);
+goal thy "(ALL F. {x}: F --> {y}:F) --> (ALL A. x:A --> y:A)";
+by (Best_tac 1);
result();
@@ -28,7 +29,7 @@
*)
(*collecting the relevant lemmas*)
-val hom_ss = ZF_ss addsimps [comp_fun,comp_fun_apply,SigmaI,apply_funtype];
+Addsimps [comp_fun, SigmaI, apply_funtype];
(*This version uses a super application of simp_tac. Needs setloop to help
proving conditions of rewrites such as comp_fun_apply;
@@ -39,7 +40,7 @@
\ (ALL x:A. ALL y:A. H`(f`<x,y>) = g`<H`x,H`y>)}) --> \
\ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \
\ (K O J) : hom(A,f,C,h)";
-by (asm_simp_tac (hom_ss setloop (K (safe_tac FOL_cs))) 1);
+by (asm_simp_tac (!simpset setloop (K (safe_tac (!claset)))) 1);
val comp_homs = result();
(*This version uses meta-level rewriting, safe_tac and asm_simp_tac*)
@@ -50,17 +51,17 @@
\ J : hom(A,f,B,g) & K : hom(B,g,C,h) --> \
\ (K O J) : hom(A,f,C,h)";
by (rewtac hom_def);
-by (safe_tac ZF_cs);
-by (asm_simp_tac hom_ss 1);
-by (asm_simp_tac hom_ss 1);
+by (safe_tac (!claset));
+by (Asm_simp_tac 1);
+by (Asm_simp_tac 1);
qed "comp_homs";
(** A characterization of functions, suggested by Tobias Nipkow **)
-goalw ZF.thy [Pi_def, function_def]
+goalw thy [Pi_def, function_def]
"r: domain(r)->B <-> r <= domain(r)*B & (ALL X. r `` (r -`` X) <= X)";
-by (best_tac ZF_cs 1);
+by (Best_tac 1);
result();
@@ -139,14 +140,14 @@
(** Yet another example... **)
-goal (merge_theories(Sum.thy,Perm.thy))
+goal Perm.thy
"(lam Z:Pow(A+B). <{x:A. Inl(x):Z}, {y:B. Inr(y):Z}>) \
\ : bij(Pow(A+B), Pow(A)*Pow(B))";
by (res_inst_tac [("d", "%<X,Y>.{Inl(x).x:X} Un {Inr(y).y:Y}")]
lam_bijective 1);
by (TRYALL (etac SigmaE));
-by (ALLGOALS (asm_simp_tac ZF_ss));
-by (ALLGOALS (fast_tac (sum_cs addSIs [equalityI])));
+by (ALLGOALS Asm_simp_tac);
+by (ALLGOALS (fast_tac (!claset addSIs [equalityI])));
val Pow_bij = result();
writeln"Reached end of file.";
--- a/src/ZF/func.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/func.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,134 +8,140 @@
(*** The Pi operator -- dependent function space ***)
-goalw ZF.thy [Pi_def]
+goalw thy [Pi_def]
"f: Pi(A,B) <-> function(f) & f<=Sigma(A,B) & A<=domain(f)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Pi_iff";
(*For upward compatibility with the former definition*)
-goalw ZF.thy [Pi_def, function_def]
+goalw thy [Pi_def, function_def]
"f: Pi(A,B) <-> f<=Sigma(A,B) & (ALL x:A. EX! y. <x,y>: f)";
-by (safe_tac ZF_cs);
-by (best_tac ZF_cs 1);
-by (best_tac ZF_cs 1);
+by (safe_tac (!claset));
+by (Best_tac 1);
+by (Best_tac 1);
by (set_mp_tac 1);
-by (deepen_tac ZF_cs 3 1);
+by (Deepen_tac 3 1);
qed "Pi_iff_old";
-goal ZF.thy "!!f. f: Pi(A,B) ==> function(f)";
+goal thy "!!f. f: Pi(A,B) ==> function(f)";
by (asm_full_simp_tac (FOL_ss addsimps [Pi_iff]) 1);
qed "fun_is_function";
(**Two "destruct" rules for Pi **)
-val [major] = goalw ZF.thy [Pi_def] "f: Pi(A,B) ==> f <= Sigma(A,B)";
+val [major] = goalw thy [Pi_def] "f: Pi(A,B) ==> f <= Sigma(A,B)";
by (rtac (major RS CollectD1 RS PowD) 1);
qed "fun_is_rel";
-goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> EX! y. <a,y>: f";
-by (fast_tac (ZF_cs addSDs [Pi_iff_old RS iffD1]) 1);
+goal thy "!!f. [| f: Pi(A,B); a:A |] ==> EX! y. <a,y>: f";
+by (fast_tac ((!claset) addSDs [Pi_iff_old RS iffD1]) 1);
qed "fun_unique_Pair";
-val prems = goalw ZF.thy [Pi_def]
+val prems = goalw thy [Pi_def]
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')";
by (simp_tac (FOL_ss addsimps prems addcongs [Sigma_cong]) 1);
qed "Pi_cong";
+(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
+ flex-flex pairs and the "Check your prover" error. Most
+ Sigmas and Pis are abbreviated as * or -> *)
+
(*Weakening one function type to another; see also Pi_type*)
-goalw ZF.thy [Pi_def] "!!f. [| f: A->B; B<=D |] ==> f: A->D";
-by (best_tac ZF_cs 1);
+goalw thy [Pi_def] "!!f. [| f: A->B; B<=D |] ==> f: A->D";
+by (Best_tac 1);
qed "fun_weaken_type";
(*Empty function spaces*)
-goalw ZF.thy [Pi_def, function_def] "Pi(0,A) = {0}";
+goalw thy [Pi_def, function_def] "Pi(0,A) = {0}";
by (fast_tac eq_cs 1);
qed "Pi_empty1";
-goalw ZF.thy [Pi_def] "!!A a. a:A ==> A->0 = 0";
+goalw thy [Pi_def] "!!A a. a:A ==> A->0 = 0";
by (fast_tac eq_cs 1);
qed "Pi_empty2";
(*The empty function*)
-goalw ZF.thy [Pi_def, function_def] "0: Pi(0,B)";
-by (fast_tac ZF_cs 1);
+goalw thy [Pi_def, function_def] "0: Pi(0,B)";
+by (Fast_tac 1);
qed "empty_fun";
(*The singleton function*)
-goalw ZF.thy [Pi_def, function_def] "{<a,b>} : {a} -> {b}";
+goalw thy [Pi_def, function_def] "{<a,b>} : {a} -> {b}";
by (fast_tac eq_cs 1);
qed "singleton_fun";
+Addsimps [empty_fun, singleton_fun];
+
(*** Function Application ***)
-goalw ZF.thy [Pi_def, function_def]
+goalw thy [Pi_def, function_def]
"!!a b f. [| <a,b>: f; <a,c>: f; f: Pi(A,B) |] ==> b=c";
-by (deepen_tac ZF_cs 3 1);
+by (Deepen_tac 3 1);
qed "apply_equality2";
-goalw ZF.thy [apply_def] "!!a b f. [| <a,b>: f; f: Pi(A,B) |] ==> f`a = b";
+goalw thy [apply_def] "!!a b f. [| <a,b>: f; f: Pi(A,B) |] ==> f`a = b";
by (rtac the_equality 1);
by (rtac apply_equality2 2);
by (REPEAT (assume_tac 1));
qed "apply_equality";
(*Applying a function outside its domain yields 0*)
-goalw ZF.thy [apply_def]
+goalw thy [apply_def]
"!!a b f. [| a ~: domain(f); f: Pi(A,B) |] ==> f`a = 0";
by (rtac the_0 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "apply_0";
-goal ZF.thy "!!f. [| f: Pi(A,B); c: f |] ==> EX x:A. c = <x,f`x>";
+goal thy "!!f. [| f: Pi(A,B); c: f |] ==> EX x:A. c = <x,f`x>";
by (forward_tac [fun_is_rel] 1);
-by (fast_tac (ZF_cs addDs [apply_equality]) 1);
+by (fast_tac ((!claset) addDs [apply_equality]) 1);
qed "Pi_memberD";
-goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> <a,f`a>: f";
+goal thy "!!f. [| f: Pi(A,B); a:A |] ==> <a,f`a>: f";
by (rtac (fun_unique_Pair RS ex1E) 1);
by (resolve_tac [apply_equality RS ssubst] 3);
by (REPEAT (assume_tac 1));
qed "apply_Pair";
(*Conclusion is flexible -- use res_inst_tac or else apply_funtype below!*)
-goal ZF.thy "!!f. [| f: Pi(A,B); a:A |] ==> f`a : B(a)";
+goal thy "!!f. [| f: Pi(A,B); a:A |] ==> f`a : B(a)";
by (rtac (fun_is_rel RS subsetD RS SigmaE2) 1);
by (REPEAT (ares_tac [apply_Pair] 1));
qed "apply_type";
(*This version is acceptable to the simplifier*)
-goal ZF.thy "!!f. [| f: A->B; a:A |] ==> f`a : B";
+goal thy "!!f. [| f: A->B; a:A |] ==> f`a : B";
by (REPEAT (ares_tac [apply_type] 1));
qed "apply_funtype";
-val [major] = goal ZF.thy
+val [major] = goal thy
"f: Pi(A,B) ==> <a,b>: f <-> a:A & f`a = b";
by (cut_facts_tac [major RS fun_is_rel] 1);
-by (fast_tac (ZF_cs addSIs [major RS apply_Pair,
+by (fast_tac ((!claset) addSIs [major RS apply_Pair,
major RSN (2,apply_equality)]) 1);
qed "apply_iff";
(*Refining one Pi type to another*)
-val pi_prem::prems = goal ZF.thy
+val pi_prem::prems = goal thy
"[| f: Pi(A,C); !!x. x:A ==> f`x : B(x) |] ==> f : Pi(A,B)";
by (cut_facts_tac [pi_prem] 1);
by (asm_full_simp_tac (FOL_ss addsimps [Pi_iff]) 1);
-by (fast_tac (ZF_cs addIs prems addSDs [pi_prem RS Pi_memberD]) 1);
+by (fast_tac ((!claset) addIs prems addSDs [pi_prem RS Pi_memberD]) 1);
qed "Pi_type";
(** Elimination of membership in a function **)
-goal ZF.thy "!!a A. [| <a,b> : f; f: Pi(A,B) |] ==> a : A";
+goal thy "!!a A. [| <a,b> : f; f: Pi(A,B) |] ==> a : A";
by (REPEAT (ares_tac [fun_is_rel RS subsetD RS SigmaD1] 1));
qed "domain_type";
-goal ZF.thy "!!b B a. [| <a,b> : f; f: Pi(A,B) |] ==> b : B(a)";
+goal thy "!!b B a. [| <a,b> : f; f: Pi(A,B) |] ==> b : B(a)";
by (etac (fun_is_rel RS subsetD RS SigmaD2) 1);
by (assume_tac 1);
qed "range_type";
-val prems = goal ZF.thy
+val prems = goal thy
"[| <a,b>: f; f: Pi(A,B); \
\ [| a:A; b:B(a); f`a = b |] ==> P \
\ |] ==> P";
@@ -146,41 +152,49 @@
(*** Lambda Abstraction ***)
-goalw ZF.thy [lam_def] "!!A b. a:A ==> <a,b(a)> : (lam x:A. b(x))";
+goalw thy [lam_def] "!!A b. a:A ==> <a,b(a)> : (lam x:A. b(x))";
by (etac RepFunI 1);
qed "lamI";
-val major::prems = goalw ZF.thy [lam_def]
+val major::prems = goalw thy [lam_def]
"[| p: (lam x:A. b(x)); !!x.[| x:A; p=<x,b(x)> |] ==> P \
\ |] ==> P";
by (rtac (major RS RepFunE) 1);
by (REPEAT (ares_tac prems 1));
qed "lamE";
-goal ZF.thy "!!a b c. [| <a,c>: (lam x:A. b(x)) |] ==> c = b(a)";
+goal thy "!!a b c. [| <a,c>: (lam x:A. b(x)) |] ==> c = b(a)";
by (REPEAT (eresolve_tac [asm_rl,lamE,Pair_inject,ssubst] 1));
qed "lamD";
-val prems = goalw ZF.thy [lam_def, Pi_def, function_def]
+val prems = goalw thy [lam_def, Pi_def, function_def]
"[| !!x. x:A ==> b(x): B(x) |] ==> (lam x:A.b(x)) : Pi(A,B)";
-by (fast_tac (ZF_cs addIs prems) 1);
+by (fast_tac ((!claset) addIs prems) 1);
qed "lam_type";
-goal ZF.thy "(lam x:A.b(x)) : A -> {b(x). x:A}";
+goal thy "(lam x:A.b(x)) : A -> {b(x). x:A}";
by (REPEAT (ares_tac [refl,lam_type,RepFunI] 1));
qed "lam_funtype";
-goal ZF.thy "!!a A. a : A ==> (lam x:A.b(x)) ` a = b(a)";
+goal thy "!!a A. a : A ==> (lam x:A.b(x)) ` a = b(a)";
by (REPEAT (ares_tac [apply_equality,lam_funtype,lamI] 1));
qed "beta";
+goalw thy [lam_def] "(lam x:0. b(x)) = 0";
+by (Simp_tac 1);
+qed "lam_empty";
+
+Addsimps [beta, lam_empty];
+
(*congruence rule for lambda abstraction*)
-val prems = goalw ZF.thy [lam_def]
+val prems = goalw thy [lam_def]
"[| A=A'; !!x. x:A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')";
by (simp_tac (FOL_ss addsimps prems addcongs [RepFun_cong]) 1);
qed "lam_cong";
-val [major] = goal ZF.thy
+Addcongs [lam_cong];
+
+val [major] = goal thy
"(!!x. x:A ==> EX! y. Q(x,y)) ==> EX f. ALL x:A. Q(x, f`x)";
by (res_inst_tac [("x", "lam x: A. THE y. Q(x,y)")] exI 1);
by (rtac ballI 1);
@@ -193,7 +207,7 @@
(** Extensionality **)
(*Semi-extensionality!*)
-val prems = goal ZF.thy
+val prems = goal thy
"[| f : Pi(A,B); g: Pi(C,D); A<=C; \
\ !!x. x:A ==> f`x = g`x |] ==> f<=g";
by (rtac subsetI 1);
@@ -203,20 +217,22 @@
by (REPEAT (ares_tac (prems@[apply_Pair,subsetD]) 1));
qed "fun_subset";
-val prems = goal ZF.thy
+val prems = goal thy
"[| f : Pi(A,B); g: Pi(A,D); \
\ !!x. x:A ==> f`x = g`x |] ==> f=g";
by (REPEAT (ares_tac (prems @ (prems RL [sym]) @
[subset_refl,equalityI,fun_subset]) 1));
qed "fun_extension";
-goal ZF.thy "!!f A B. f : Pi(A,B) ==> (lam x:A. f`x) = f";
+goal thy "!!f A B. f : Pi(A,B) ==> (lam x:A. f`x) = f";
by (rtac fun_extension 1);
by (REPEAT (ares_tac [lam_type,apply_type,beta] 1));
qed "eta";
+Addsimps [eta];
+
(*Every element of Pi(A,B) may be expressed as a lambda abstraction!*)
-val prems = goal ZF.thy
+val prems = goal thy
"[| f: Pi(A,B); \
\ !!b. [| ALL x:A. b(x):B(x); f = (lam x:A.b(x)) |] ==> P \
\ |] ==> P";
@@ -228,11 +244,11 @@
(** Images of functions **)
-goalw ZF.thy [lam_def] "!!C A. C <= A ==> (lam x:A.b(x)) `` C = {b(x). x:C}";
+goalw thy [lam_def] "!!C A. C <= A ==> (lam x:A.b(x)) `` C = {b(x). x:C}";
by (fast_tac eq_cs 1);
qed "image_lam";
-goal ZF.thy "!!C A. [| f : Pi(A,B); C <= A |] ==> f``C = {f`x. x:C}";
+goal thy "!!C A. [| f : Pi(A,B); C <= A |] ==> f``C = {f`x. x:C}";
by (etac (eta RS subst) 1);
by (asm_full_simp_tac (FOL_ss addsimps [beta, image_lam, subset_iff]
addcongs [RepFun_cong]) 1);
@@ -241,38 +257,44 @@
(*** properties of "restrict" ***)
-goalw ZF.thy [restrict_def,lam_def]
+goalw thy [restrict_def,lam_def]
"!!f A. [| f: Pi(C,B); A<=C |] ==> restrict(f,A) <= f";
-by (fast_tac (ZF_cs addIs [apply_Pair]) 1);
+by (fast_tac ((!claset) addIs [apply_Pair]) 1);
qed "restrict_subset";
-val prems = goalw ZF.thy [restrict_def]
+val prems = goalw thy [restrict_def]
"[| !!x. x:A ==> f`x: B(x) |] ==> restrict(f,A) : Pi(A,B)";
by (rtac lam_type 1);
by (eresolve_tac prems 1);
qed "restrict_type";
-val [pi,subs] = goal ZF.thy
+val [pi,subs] = goal thy
"[| f: Pi(C,B); A<=C |] ==> restrict(f,A) : Pi(A,B)";
by (rtac (pi RS apply_type RS restrict_type) 1);
by (etac (subs RS subsetD) 1);
qed "restrict_type2";
-goalw ZF.thy [restrict_def] "!!a A. a : A ==> restrict(f,A) ` a = f`a";
+goalw thy [restrict_def] "!!a A. a : A ==> restrict(f,A) ` a = f`a";
by (etac beta 1);
qed "restrict";
+goalw thy [restrict_def] "restrict(f,0) = 0";
+by (Simp_tac 1);
+qed "restrict_empty";
+
+Addsimps [restrict, restrict_empty];
+
(*NOT SAFE as a congruence rule for the simplifier! Can cause it to fail!*)
-val prems = goalw ZF.thy [restrict_def]
+val prems = goalw thy [restrict_def]
"[| A=B; !!x. x:B ==> f`x=g`x |] ==> restrict(f,A) = restrict(g,B)";
by (REPEAT (ares_tac (prems@[lam_cong]) 1));
qed "restrict_eqI";
-goalw ZF.thy [restrict_def, lam_def] "domain(restrict(f,C)) = C";
+goalw thy [restrict_def, lam_def] "domain(restrict(f,C)) = C";
by (fast_tac eq_cs 1);
qed "domain_restrict";
-val [prem] = goalw ZF.thy [restrict_def]
+val [prem] = goalw thy [restrict_def]
"A<=C ==> restrict(lam x:C. b(x), A) = (lam x:A. b(x))";
by (rtac (refl RS lam_cong) 1);
by (etac (prem RS subsetD RS beta) 1); (*easier than calling simp_tac*)
@@ -284,14 +306,14 @@
(** The Union of a set of COMPATIBLE functions is a function **)
-goalw ZF.thy [function_def]
+goalw thy [function_def]
"!!S. [| ALL x:S. function(x); \
\ ALL x:S. ALL y:S. x<=y | y<=x |] ==> \
\ function(Union(S))";
-by (fast_tac (ZF_cs addSDs [bspec RS bspec]) 1);
+by (fast_tac ((!claset) addSDs [bspec RS bspec]) 1);
qed "function_Union";
-goalw ZF.thy [Pi_def]
+goalw thy [Pi_def]
"!!S. [| ALL f:S. EX C D. f:C->D; \
\ ALL f:S. ALL y:S. f<=y | y<=f |] ==> \
\ Union(S) : domain(Union(S)) -> range(Union(S))";
@@ -305,30 +327,30 @@
Un_upper1 RSN (2, subset_trans),
Un_upper2 RSN (2, subset_trans)];
-goal ZF.thy
+goal thy
"!!f. [| f: A->B; g: C->D; A Int C = 0 |] ==> \
\ (f Un g) : (A Un C) -> (B Un D)";
(*Solve the product and domain subgoals using distributive laws*)
by (asm_full_simp_tac (FOL_ss addsimps [Pi_iff,extension]@Un_rls) 1);
by (asm_simp_tac (FOL_ss addsimps [function_def]) 1);
-by (safe_tac ZF_cs);
+by (safe_tac (!claset));
(*Solve the two cases that contradict A Int C = 0*)
-by (deepen_tac ZF_cs 2 2);
-by (deepen_tac ZF_cs 2 2);
+by (Deepen_tac 2 2);
+by (Deepen_tac 2 2);
by (rewtac function_def);
by (REPEAT_FIRST (dtac (spec RS spec)));
-by (deepen_tac ZF_cs 1 1);
-by (deepen_tac ZF_cs 1 1);
+by (Deepen_tac 1 1);
+by (Deepen_tac 1 1);
qed "fun_disjoint_Un";
-goal ZF.thy
+goal thy
"!!f g a. [| a:A; f: A->B; g: C->D; A Int C = 0 |] ==> \
\ (f Un g)`a = f`a";
by (rtac (apply_Pair RS UnI1 RS apply_equality) 1);
by (REPEAT (ares_tac [fun_disjoint_Un] 1));
qed "fun_disjoint_apply1";
-goal ZF.thy
+goal thy
"!!f g c. [| c:C; f: A->B; g: C->D; A Int C = 0 |] ==> \
\ (f Un g)`c = g`c";
by (rtac (apply_Pair RS UnI2 RS apply_equality) 1);
@@ -337,61 +359,64 @@
(** Domain and range of a function/relation **)
-goalw ZF.thy [Pi_def] "!!f. f : Pi(A,B) ==> domain(f)=A";
+goalw thy [Pi_def] "!!f. f : Pi(A,B) ==> domain(f)=A";
by (fast_tac eq_cs 1);
qed "domain_of_fun";
-goal ZF.thy "!!f. [| f : Pi(A,B); a: A |] ==> f`a : range(f)";
+goal thy "!!f. [| f : Pi(A,B); a: A |] ==> f`a : range(f)";
by (etac (apply_Pair RS rangeI) 1);
by (assume_tac 1);
qed "apply_rangeI";
-val [major] = goal ZF.thy "f : Pi(A,B) ==> f : A->range(f)";
+val [major] = goal thy "f : Pi(A,B) ==> f : A->range(f)";
by (rtac (major RS Pi_type) 1);
by (etac (major RS apply_rangeI) 1);
qed "range_of_fun";
(*** Extensions of functions ***)
-goal ZF.thy
+goal thy
"!!f A B. [| f: A->B; c~:A |] ==> cons(<c,b>,f) : cons(c,A) -> cons(b,B)";
by (forward_tac [singleton_fun RS fun_disjoint_Un] 1);
by (asm_full_simp_tac (FOL_ss addsimps [cons_eq]) 2);
by (fast_tac eq_cs 1);
qed "fun_extend";
-goal ZF.thy
+goal thy
"!!f A B. [| f: A->B; c~:A; b: B |] ==> cons(<c,b>,f) : cons(c,A) -> B";
-by (fast_tac (ZF_cs addEs [fun_extend RS fun_weaken_type]) 1);
+by (fast_tac ((!claset) addEs [fun_extend RS fun_weaken_type]) 1);
qed "fun_extend3";
-goal ZF.thy "!!f A B. [| f: A->B; a:A; c~:A |] ==> cons(<c,b>,f)`a = f`a";
+goal thy "!!f A B. [| f: A->B; a:A; c~:A |] ==> cons(<c,b>,f)`a = f`a";
by (rtac (apply_Pair RS consI2 RS apply_equality) 1);
by (rtac fun_extend 3);
by (REPEAT (assume_tac 1));
qed "fun_extend_apply1";
-goal ZF.thy "!!f A B. [| f: A->B; c~:A |] ==> cons(<c,b>,f)`c = b";
+goal thy "!!f A B. [| f: A->B; c~:A |] ==> cons(<c,b>,f)`c = b";
by (rtac (consI1 RS apply_equality) 1);
by (rtac fun_extend 1);
by (REPEAT (assume_tac 1));
qed "fun_extend_apply2";
+bind_thm ("singleton_apply", [singletonI, singleton_fun] MRS apply_equality);
+Addsimps [singleton_apply];
+
(*For Finite.ML. Inclusion of right into left is easy*)
-goal ZF.thy
+goal thy
"!!c. c ~: A ==> cons(c,A) -> B = (UN f: A->B. UN b:B. {cons(<c,b>, f)})";
by (rtac equalityI 1);
-by (safe_tac (ZF_cs addSEs [fun_extend3]));
+by (safe_tac ((!claset) addSEs [fun_extend3]));
(*Inclusion of left into right*)
by (subgoal_tac "restrict(x, A) : A -> B" 1);
-by (fast_tac (ZF_cs addEs [restrict_type2]) 2);
+by (fast_tac ((!claset) addEs [restrict_type2]) 2);
by (rtac UN_I 1 THEN assume_tac 1);
by (rtac (apply_funtype RS UN_I) 1 THEN REPEAT (ares_tac [consI1] 1));
-by (simp_tac (FOL_ss addsimps cons_iff::mem_simps) 1);
+by (Simp_tac 1);
by (rtac fun_extension 1 THEN REPEAT (ares_tac [fun_extend] 1));
by (etac consE 1);
by (ALLGOALS
- (asm_simp_tac (FOL_ss addsimps [restrict, fun_extend_apply1,
- fun_extend_apply2])));
+ (asm_simp_tac (!simpset addsimps [restrict, fun_extend_apply1,
+ fun_extend_apply2])));
qed "cons_fun_eq";
--- a/src/ZF/func.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/func.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,3 @@
(*Dummy theory to document dependencies *)
-func = "domrange" + "equalities"
+func = domrange + equalities
--- a/src/ZF/mono.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/mono.ML Fri Jan 03 15:01:55 1997 +0100
@@ -6,71 +6,73 @@
Monotonicity of various operations (for lattice properties see subset.ML)
*)
+open mono;
+
(** Replacement, in its various formulations **)
(*Not easy to express monotonicity in P, since any "bigger" predicate
would have to be single-valued*)
-goal ZF.thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)";
-by (fast_tac (ZF_cs addSEs [ReplaceE]) 1);
+goal thy "!!A B. A<=B ==> Replace(A,P) <= Replace(B,P)";
+by (fast_tac (!claset addSEs [ReplaceE]) 1);
qed "Replace_mono";
-goal ZF.thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. A<=B ==> {f(x). x:A} <= {f(x). x:B}";
+by (Fast_tac 1);
qed "RepFun_mono";
-goal ZF.thy "!!A B. A<=B ==> Pow(A) <= Pow(B)";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. A<=B ==> Pow(A) <= Pow(B)";
+by (Fast_tac 1);
qed "Pow_mono";
-goal ZF.thy "!!A B. A<=B ==> Union(A) <= Union(B)";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. A<=B ==> Union(A) <= Union(B)";
+by (Fast_tac 1);
qed "Union_mono";
-val prems = goal ZF.thy
+val prems = goal thy
"[| A<=C; !!x. x:A ==> B(x)<=D(x) \
\ |] ==> (UN x:A. B(x)) <= (UN x:C. D(x))";
-by (fast_tac (ZF_cs addIs (prems RL [subsetD])) 1);
+by (fast_tac (!claset addIs (prems RL [subsetD])) 1);
qed "UN_mono";
(*Intersection is ANTI-monotonic. There are TWO premises! *)
-goal ZF.thy "!!A B. [| A<=B; a:A |] ==> Inter(B) <= Inter(A)";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B. [| A<=B; a:A |] ==> Inter(B) <= Inter(A)";
+by (Fast_tac 1);
qed "Inter_anti_mono";
-goal ZF.thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)";
-by (fast_tac ZF_cs 1);
+goal thy "!!C D. C<=D ==> cons(a,C) <= cons(a,D)";
+by (Fast_tac 1);
qed "cons_mono";
-goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Un B <= C Un D";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B C D. [| A<=C; B<=D |] ==> A Un B <= C Un D";
+by (Fast_tac 1);
qed "Un_mono";
-goal ZF.thy "!!A B C D. [| A<=C; B<=D |] ==> A Int B <= C Int D";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B C D. [| A<=C; B<=D |] ==> A Int B <= C Int D";
+by (Fast_tac 1);
qed "Int_mono";
-goal ZF.thy "!!A B C D. [| A<=C; D<=B |] ==> A-B <= C-D";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B C D. [| A<=C; D<=B |] ==> A-B <= C-D";
+by (Fast_tac 1);
qed "Diff_mono";
(** Standard products, sums and function spaces **)
-goal ZF.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \
+goal thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \
\ Sigma(A,B) <= Sigma(C,D)";
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "Sigma_mono_lemma";
val Sigma_mono = ballI RSN (2,Sigma_mono_lemma);
-goalw Sum.thy sum_defs "!!A B C D. [| A<=C; B<=D |] ==> A+B <= C+D";
+goalw thy sum_defs "!!A B C D. [| A<=C; B<=D |] ==> A+B <= C+D";
by (REPEAT (ares_tac [subset_refl,Un_mono,Sigma_mono] 1));
qed "sum_mono";
(*Note that B->A and C->A are typically disjoint!*)
-goal ZF.thy "!!A B C. B<=C ==> A->B <= A->C";
-by (fast_tac (ZF_cs addIs [lam_type] addEs [Pi_lamE]) 1);
+goal thy "!!A B C. B<=C ==> A->B <= A->C";
+by (fast_tac (!claset addIs [lam_type] addEs [Pi_lamE]) 1);
qed "Pi_mono";
-goalw ZF.thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)";
+goalw thy [lam_def] "!!A B. A<=B ==> Lambda(A,c) <= Lambda(B,c)";
by (etac RepFun_mono 1);
qed "lam_mono";
@@ -82,7 +84,7 @@
goal QPair.thy "!!A B C D. [| A<=C; ALL x:A. B(x) <= D(x) |] ==> \
\ QSigma(A,B) <= QSigma(C,D)";
-by (fast_tac qpair_cs 1);
+by (Fast_tac 1);
qed "QSigma_mono_lemma";
val QSigma_mono = ballI RSN (2,QSigma_mono_lemma);
@@ -95,67 +97,67 @@
qed "QInr_mono";
goal QPair.thy "!!A B C D. [| A<=C; B<=D |] ==> A <+> B <= C <+> D";
-by (fast_tac qsum_cs 1);
+by (Fast_tac 1);
qed "qsum_mono";
(** Converse, domain, range, field **)
-goal ZF.thy "!!r s. r<=s ==> converse(r) <= converse(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> converse(r) <= converse(s)";
+by (Fast_tac 1);
qed "converse_mono";
-goal ZF.thy "!!r s. r<=s ==> domain(r)<=domain(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> domain(r)<=domain(s)";
+by (Fast_tac 1);
qed "domain_mono";
bind_thm ("domain_rel_subset", [domain_mono, domain_subset] MRS subset_trans);
-goal ZF.thy "!!r s. r<=s ==> range(r)<=range(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> range(r)<=range(s)";
+by (Fast_tac 1);
qed "range_mono";
bind_thm ("range_rel_subset", [range_mono, range_subset] MRS subset_trans);
-goal ZF.thy "!!r s. r<=s ==> field(r)<=field(s)";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. r<=s ==> field(r)<=field(s)";
+by (Fast_tac 1);
qed "field_mono";
-goal ZF.thy "!!r A. r <= A*A ==> field(r) <= A";
+goal thy "!!r A. r <= A*A ==> field(r) <= A";
by (etac (field_mono RS subset_trans) 1);
-by (fast_tac ZF_cs 1);
+by (Fast_tac 1);
qed "field_rel_subset";
(** Images **)
-val [prem1,prem2] = goal ZF.thy
+val [prem1,prem2] = goal thy
"[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r``A <= s``B";
-by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1);
+by (fast_tac (!claset addIs [prem1, prem2 RS subsetD]) 1);
qed "image_pair_mono";
-val [prem1,prem2] = goal ZF.thy
+val [prem1,prem2] = goal thy
"[| !! x y. <x,y>:r ==> <x,y>:s; A<=B |] ==> r-``A <= s-``B";
-by (fast_tac (ZF_cs addIs [prem1, prem2 RS subsetD]) 1);
+by (fast_tac (!claset addIs [prem1, prem2 RS subsetD]) 1);
qed "vimage_pair_mono";
-goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r``A <= s``B";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. [| r<=s; A<=B |] ==> r``A <= s``B";
+by (Fast_tac 1);
qed "image_mono";
-goal ZF.thy "!!r s. [| r<=s; A<=B |] ==> r-``A <= s-``B";
-by (fast_tac ZF_cs 1);
+goal thy "!!r s. [| r<=s; A<=B |] ==> r-``A <= s-``B";
+by (Fast_tac 1);
qed "vimage_mono";
-val [sub,PQimp] = goal ZF.thy
+val [sub,PQimp] = goal thy
"[| A<=B; !!x. x:A ==> P(x) --> Q(x) |] ==> Collect(A,P) <= Collect(B,Q)";
-by (fast_tac (ZF_cs addIs [sub RS subsetD, PQimp RS mp]) 1);
+by (fast_tac (!claset addIs [sub RS subsetD, PQimp RS mp]) 1);
qed "Collect_mono";
(** Monotonicity of implications -- some could go to FOL **)
-goal ZF.thy "!!A B x. A<=B ==> x:A --> x:B";
-by (fast_tac ZF_cs 1);
+goal thy "!!A B x. A<=B ==> x:A --> x:B";
+by (Fast_tac 1);
qed "in_mono";
goal IFOL.thy "!!P1 P2 Q1 Q2. [| P1-->Q1; P2-->Q2 |] ==> (P1&P2) --> (Q1&Q2)";
@@ -177,12 +179,12 @@
val [PQimp] = goal IFOL.thy
"[| !!x. P(x) --> Q(x) |] ==> (EX x.P(x)) --> (EX x.Q(x))";
-by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1);
+by (fast_tac (!claset addIs [PQimp RS mp]) 1);
qed "ex_mono";
val [PQimp] = goal IFOL.thy
"[| !!x. P(x) --> Q(x) |] ==> (ALL x.P(x)) --> (ALL x.Q(x))";
-by (fast_tac (FOL_cs addIs [PQimp RS mp]) 1);
+by (fast_tac (!claset addIs [PQimp RS mp]) 1);
qed "all_mono";
(*Used in intr_elim.ML and in individual datatype definitions*)
--- a/src/ZF/mono.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/mono.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,4 @@
(*Dummy theory to document dependencies *)
-mono = QPair + Sum
+mono = QPair + Sum + domrange
+
--- a/src/ZF/pair.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/pair.ML Fri Jan 03 15:01:55 1997 +0100
@@ -8,43 +8,44 @@
(** Lemmas for showing that <a,b> uniquely determines a and b **)
-qed_goal "singleton_eq_iff" ZF.thy
+qed_goal "singleton_eq_iff" thy
"{a} = {b} <-> a=b"
(fn _=> [ (resolve_tac [extension RS iff_trans] 1),
- (fast_tac upair_cs 1) ]);
+ (Fast_tac 1) ]);
-qed_goal "doubleton_eq_iff" ZF.thy
+qed_goal "doubleton_eq_iff" thy
"{a,b} = {c,d} <-> (a=c & b=d) | (a=d & b=c)"
(fn _=> [ (resolve_tac [extension RS iff_trans] 1),
- (fast_tac upair_cs 1) ]);
+ (Fast_tac 1) ]);
-qed_goalw "Pair_iff" ZF.thy [Pair_def]
+qed_goalw "Pair_iff" thy [Pair_def]
"<a,b> = <c,d> <-> a=c & b=d"
- (fn _=> [ (simp_tac (FOL_ss addsimps [doubleton_eq_iff]) 1),
- (fast_tac FOL_cs 1) ]);
+ (fn _=> [ (simp_tac (!simpset addsimps [doubleton_eq_iff]) 1),
+ (Fast_tac 1) ]);
-bind_thm ("Pair_inject", (Pair_iff RS iffD1 RS conjE));
+Addsimps [Pair_iff];
+
+bind_thm ("Pair_inject", Pair_iff RS iffD1 RS conjE);
-qed_goal "Pair_inject1" ZF.thy "<a,b> = <c,d> ==> a=c"
- (fn [major]=>
- [ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]);
+AddSEs [Pair_inject];
+
+bind_thm ("Pair_inject1", Pair_iff RS iffD1 RS conjunct1);
+bind_thm ("Pair_inject2", Pair_iff RS iffD1 RS conjunct2);
-qed_goal "Pair_inject2" ZF.thy "<a,b> = <c,d> ==> b=d"
- (fn [major]=>
- [ (rtac (major RS Pair_inject) 1), (assume_tac 1) ]);
+qed_goalw "Pair_not_0" thy [Pair_def] "<a,b> ~= 0"
+ (fn _ => [ (fast_tac (!claset addEs [equalityE]) 1) ]);
-qed_goalw "Pair_neq_0" ZF.thy [Pair_def] "<a,b>=0 ==> P"
- (fn [major]=>
- [ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1),
- (rtac consI1 1) ]);
+bind_thm ("Pair_neq_0", Pair_not_0 RS notE);
-qed_goalw "Pair_neq_fst" ZF.thy [Pair_def] "<a,b>=a ==> P"
+AddSEs [Pair_neq_0, sym RS Pair_neq_0];
+
+qed_goalw "Pair_neq_fst" thy [Pair_def] "<a,b>=a ==> P"
(fn [major]=>
[ (rtac (consI1 RS mem_asym RS FalseE) 1),
(rtac (major RS subst) 1),
(rtac consI1 1) ]);
-qed_goalw "Pair_neq_snd" ZF.thy [Pair_def] "<a,b>=b ==> P"
+qed_goalw "Pair_neq_snd" thy [Pair_def] "<a,b>=b ==> P"
(fn [major]=>
[ (rtac (consI1 RS consI2 RS mem_asym RS FalseE) 1),
(rtac (major RS subst) 1),
@@ -54,12 +55,20 @@
(*** Sigma: Disjoint union of a family of sets
Generalizes Cartesian product ***)
-qed_goalw "SigmaI" ZF.thy [Sigma_def]
- "[| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)"
- (fn prems=> [ (REPEAT (resolve_tac (prems@[singletonI,UN_I]) 1)) ]);
+qed_goalw "Sigma_iff" thy [Sigma_def] "<a,b>: Sigma(A,B) <-> a:A & b:B(a)"
+ (fn _ => [ Fast_tac 1 ]);
+
+Addsimps [Sigma_iff];
+
+qed_goal "SigmaI" thy
+ "!!a b. [| a:A; b:B(a) |] ==> <a,b> : Sigma(A,B)"
+ (fn _ => [ Asm_simp_tac 1 ]);
+
+bind_thm ("SigmaD1", Sigma_iff RS iffD1 RS conjunct1);
+bind_thm ("SigmaD2", Sigma_iff RS iffD1 RS conjunct2);
(*The general elimination rule*)
-qed_goalw "SigmaE" ZF.thy [Sigma_def]
+qed_goalw "SigmaE" thy [Sigma_def]
"[| c: Sigma(A,B); \
\ !!x y.[| x:A; y:B(x); c=<x,y> |] ==> P \
\ |] ==> P"
@@ -67,22 +76,7 @@
[ (cut_facts_tac [major] 1),
(REPEAT (eresolve_tac [UN_E, singletonE] 1 ORELSE ares_tac prems 1)) ]);
-(** Elimination of <a,b>:A*B -- introduces no eigenvariables **)
-qed_goal "SigmaD1" ZF.thy "<a,b> : Sigma(A,B) ==> a : A"
- (fn [major]=>
- [ (rtac (major RS SigmaE) 1),
- (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
-
-qed_goal "SigmaD2" ZF.thy "<a,b> : Sigma(A,B) ==> b : B(a)"
- (fn [major]=>
- [ (rtac (major RS SigmaE) 1),
- (REPEAT (eresolve_tac [asm_rl,Pair_inject,ssubst] 1)) ]);
-
-(*Also provable via
- rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac)
- THEN prune_params_tac)
- (read_instantiate [("c","<a,b>")] SigmaE); *)
-qed_goal "SigmaE2" ZF.thy
+qed_goal "SigmaE2" thy
"[| <a,b> : Sigma(A,B); \
\ [| a:A; b:B(a) |] ==> P \
\ |] ==> P"
@@ -91,136 +85,92 @@
(rtac (major RS SigmaD1) 1),
(rtac (major RS SigmaD2) 1) ]);
-qed_goalw "Sigma_cong" ZF.thy [Sigma_def]
+qed_goalw "Sigma_cong" thy [Sigma_def]
"[| A=A'; !!x. x:A' ==> B(x)=B'(x) |] ==> \
\ Sigma(A,B) = Sigma(A',B')"
- (fn prems=> [ (simp_tac (FOL_ss addsimps prems addcongs [RepFun_cong]) 1) ]);
+ (fn prems=> [ (simp_tac (!simpset addsimps prems) 1) ]);
+
-qed_goal "Sigma_empty1" ZF.thy "Sigma(0,B) = 0"
- (fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]);
+(*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
+ flex-flex pairs and the "Check your prover" error. Most
+ Sigmas and Pis are abbreviated as * or -> *)
-qed_goal "Sigma_empty2" ZF.thy "A*0 = 0"
- (fn _ => [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [SigmaE]) 1) ]);
+AddSIs [SigmaI];
+AddSEs [SigmaE2, SigmaE];
-val pair_cs = upair_cs
- addSIs [SigmaI]
- addSEs [SigmaE2, SigmaE, Pair_inject, make_elim succ_inject,
- Pair_neq_0, sym RS Pair_neq_0, succ_neq_0, sym RS succ_neq_0];
+qed_goal "Sigma_empty1" thy "Sigma(0,B) = 0"
+ (fn _ => [ (fast_tac (!claset addIs [equalityI]) 1) ]);
+
+qed_goal "Sigma_empty2" thy "A*0 = 0"
+ (fn _ => [ (fast_tac (!claset addIs [equalityI]) 1) ]);
+
+Addsimps [Sigma_empty1, Sigma_empty2];
(*** Projections: fst, snd ***)
-qed_goalw "fst_conv" ZF.thy [fst_def] "fst(<a,b>) = a"
- (fn _=>
- [ (fast_tac (pair_cs addIs [the_equality]) 1) ]);
+qed_goalw "fst_conv" thy [fst_def] "fst(<a,b>) = a"
+ (fn _=> [ (fast_tac (!claset addIs [the_equality]) 1) ]);
-qed_goalw "snd_conv" ZF.thy [snd_def] "snd(<a,b>) = b"
- (fn _=>
- [ (fast_tac (pair_cs addIs [the_equality]) 1) ]);
+qed_goalw "snd_conv" thy [snd_def] "snd(<a,b>) = b"
+ (fn _=> [ (fast_tac (!claset addIs [the_equality]) 1) ]);
-val pair_ss = FOL_ss addsimps [fst_conv,snd_conv];
+Addsimps [fst_conv,snd_conv];
-qed_goal "fst_type" ZF.thy
- "!!p. p:Sigma(A,B) ==> fst(p) : A"
- (fn _=> [ (fast_tac (pair_cs addss pair_ss) 1) ]);
+qed_goal "fst_type" thy "!!p. p:Sigma(A,B) ==> fst(p) : A"
+ (fn _=> [ Auto_tac() ]);
-qed_goal "snd_type" ZF.thy
- "!!p. p:Sigma(A,B) ==> snd(p) : B(fst(p))"
- (fn _=> [ (fast_tac (pair_cs addss pair_ss) 1) ]);
+qed_goal "snd_type" thy "!!p. p:Sigma(A,B) ==> snd(p) : B(fst(p))"
+ (fn _=> [ Auto_tac() ]);
-goal ZF.thy "!!a A B. a: Sigma(A,B) ==> <fst(a),snd(a)> = a";
-by (etac SigmaE 1);
-by (asm_simp_tac pair_ss 1);
-qed "Pair_fst_snd_eq";
+qed_goal "Pair_fst_snd_eq" thy
+ "!!a A B. a: Sigma(A,B) ==> <fst(a),snd(a)> = a"
+ (fn _=> [ Auto_tac() ]);
(*** Eliminator - split ***)
(*A META-equality, so that it applies to higher types as well...*)
-qed_goalw "split" ZF.thy [split_def]
- "split(%x y.c(x,y), <a,b>) == c(a,b)"
- (fn _ => [ (simp_tac pair_ss 1),
+qed_goalw "split" thy [split_def] "split(%x y.c(x,y), <a,b>) == c(a,b)"
+ (fn _ => [ (Simp_tac 1),
(rtac reflexive_thm 1) ]);
-qed_goal "split_type" ZF.thy
+Addsimps [split];
+
+qed_goal "split_type" thy
"[| p:Sigma(A,B); \
\ !!x y.[| x:A; y:B(x) |] ==> c(x,y):C(<x,y>) \
\ |] ==> split(%x y.c(x,y), p) : C(p)"
(fn major::prems=>
[ (rtac (major RS SigmaE) 1),
- (asm_simp_tac (pair_ss addsimps (split::prems)) 1) ]);
+ (asm_simp_tac (!simpset addsimps prems) 1) ]);
-goalw ZF.thy [split_def]
+goalw thy [split_def]
"!!u. u: A*B ==> \
\ R(split(c,u)) <-> (ALL x:A. ALL y:B. u = <x,y> --> R(c(x,y)))";
-by (etac SigmaE 1);
-by (asm_simp_tac pair_ss 1);
-by (fast_tac pair_cs 1);
+by (Auto_tac());
qed "expand_split";
(*** split for predicates: result type o ***)
-goalw ZF.thy [split_def] "!!R a b. R(a,b) ==> split(R, <a,b>)";
-by (asm_simp_tac pair_ss 1);
+goalw thy [split_def] "!!R a b. R(a,b) ==> split(R, <a,b>)";
+by (Asm_simp_tac 1);
qed "splitI";
-val major::sigma::prems = goalw ZF.thy [split_def]
+val major::sigma::prems = goalw thy [split_def]
"[| split(R,z); z:Sigma(A,B); \
\ !!x y. [| z = <x,y>; R(x,y) |] ==> P \
\ |] ==> P";
by (rtac (sigma RS SigmaE) 1);
by (cut_facts_tac [major] 1);
-by (asm_full_simp_tac (pair_ss addsimps prems) 1);
+by (REPEAT (ares_tac prems 1));
+by (Asm_full_simp_tac 1);
qed "splitE";
-goalw ZF.thy [split_def] "!!R a b. split(R,<a,b>) ==> R(a,b)";
-by (asm_full_simp_tac pair_ss 1);
+goalw thy [split_def] "!!R a b. split(R,<a,b>) ==> R(a,b)";
+by (Full_simp_tac 1);
qed "splitD";
-(*** Basic simplification for ZF; see simpdata.ML for full version ***)
-fun prove_fun s =
- (writeln s; prove_goal ZF.thy s
- (fn prems => [ (cut_facts_tac prems 1),
- (fast_tac (pair_cs addSIs [equalityI]) 1) ]));
-
-(*INCLUDED IN ZF_ss*)
-val mem_simps = map prove_fun
- [ "a : 0 <-> False",
- "a : A Un B <-> a:A | a:B",
- "a : A Int B <-> a:A & a:B",
- "a : A-B <-> a:A & ~a:B",
- "<a,b>: Sigma(A,B) <-> a:A & b:B(a)",
- "a : Collect(A,P) <-> a:A & P(a)" ];
-
-goal ZF.thy "{x.x:A} = A";
-by (fast_tac (pair_cs addSIs [equalityI]) 1);
-qed "triv_RepFun";
-
-(*INCLUDED: should be? And what about cons(a,b)?*)
-val bquant_simps = map prove_fun
- [ "(ALL x:0.P(x)) <-> True",
- "(EX x:0.P(x)) <-> False",
- "(ALL x:succ(i).P(x)) <-> P(i) & (ALL x:i.P(x))",
- "(EX x:succ(i).P(x)) <-> P(i) | (EX x:i.P(x))",
- "(ALL x:cons(a,B).P(x)) <-> P(a) & (ALL x:B.P(x))",
- "(EX x:cons(a,B).P(x)) <-> P(a) | (EX x:B.P(x))" ];
-
-val Collect_simps = map prove_fun
- [ "{x:0. P(x)} = 0",
- "{x:A. False} = 0",
- "{x:A. True} = A",
- "RepFun(0,f) = 0",
- "RepFun(succ(i),f) = cons(f(i), RepFun(i,f))",
- "RepFun(cons(a,B),f) = cons(f(a), RepFun(B,f))" ];
-
-val UnInt_simps = map prove_fun
- [ "0 Un A = A", "A Un 0 = A",
- "0 Int A = 0", "A Int 0 = 0",
- "0-A = 0", "A-0 = A",
- "Union(0) = 0",
- "Union(cons(b,A)) = b Un Union(A)",
- "Inter({b}) = b"];
-
--- a/src/ZF/pair.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/pair.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,4 @@
(*Dummy theory to document dependencies *)
-pair = "upair"
+pair = upair
+
--- a/src/ZF/simpdata.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/simpdata.ML Fri Jan 03 15:01:55 1997 +0100
@@ -3,35 +3,43 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
Copyright 1991 University of Cambridge
-Rewriting for ZF set theory -- based on FOL rewriting
+Rewriting for ZF set theory: specialized extraction of rewrites from theorems
*)
-(** Tactics for type checking -- from CTT **)
+(** Rewriting **)
-fun is_rigid_elem (Const("Trueprop",_) $ (Const("op :",_) $ a $ _)) =
- not (is_Var (head_of a))
- | is_rigid_elem _ = false;
-
-(*Try solving a:A by assumption provided a is rigid!*)
-val test_assume_tac = SUBGOAL(fn (prem,i) =>
- if is_rigid_elem (Logic.strip_assums_concl prem)
- then assume_tac i else eq_assume_tac i);
+(*For proving rewrite rules*)
+fun prove_fun s =
+ (writeln s; prove_goal thy s
+ (fn prems => [ (cut_facts_tac prems 1),
+ (fast_tac (!claset addSIs [equalityI]) 1) ]));
-(*Type checking solves a:?A (a rigid, ?A maybe flexible).
- match_tac is too strict; would refuse to instantiate ?A*)
-fun typechk_step_tac tyrls =
- FIRSTGOAL (test_assume_tac ORELSE' filt_resolve_tac tyrls 3);
-
-fun typechk_tac tyrls = REPEAT (typechk_step_tac tyrls);
+(*Are all these really suitable?*)
+Addsimps (map prove_fun
+ ["(ALL x:0.P(x)) <-> True",
+ "(EX x:0.P(x)) <-> False",
+ "(ALL x:succ(i).P(x)) <-> P(i) & (ALL x:i.P(x))",
+ "(EX x:succ(i).P(x)) <-> P(i) | (EX x:i.P(x))",
+ "(ALL x:cons(a,B).P(x)) <-> P(a) & (ALL x:B.P(x))",
+ "(EX x:cons(a,B).P(x)) <-> P(a) | (EX x:B.P(x))",
+ "(ALL x: RepFun(A,f). P(x)) <-> (ALL y:A. P(f(y)))",
+ "(EX x: RepFun(A,f). P(x)) <-> (EX y:A. P(f(y)))"]);
-val ZF_typechecks =
- [if_type, lam_type, SigmaI, apply_type, split_type, consI1];
+Addsimps (map prove_fun
+ ["{x:0. P(x)} = 0",
+ "{x:A. False} = 0",
+ "{x:A. True} = A",
+ "RepFun(0,f) = 0",
+ "RepFun(succ(i),f) = cons(f(i), RepFun(i,f))",
+ "RepFun(cons(a,B),f) = cons(f(a), RepFun(B,f))"]);
-(*Instantiates variables in typing conditions.
- drawback: does not simplify conjunctions*)
-fun type_auto_tac tyrls hyps = SELECT_GOAL
- (DEPTH_SOLVE (typechk_step_tac (tyrls@hyps)
- ORELSE ares_tac [TrueI,refl,iff_refl,ballI,allI,conjI,impI] 1));
+Addsimps (map prove_fun
+ ["0 Un A = A", "A Un 0 = A",
+ "0 Int A = 0", "A Int 0 = 0",
+ "0-A = 0", "A-0 = A",
+ "Union(0) = 0",
+ "Union(cons(b,A)) = b Un Union(A)",
+ "Inter({b}) = b"]);
(** New version of mk_rew_rules **)
@@ -70,24 +78,8 @@
("op -", [DiffD1,DiffD2]),
("op Int", [IntD1,IntD2])];
-val ZF_simps =
- [empty_subsetI, consI1, cons_not_0, cons_not_0 RS not_sym,
- succI1, succ_not_0, succ_not_0 RS not_sym, succ_inject_iff,
- ball_simp, if_true, if_false,
- beta, eta, restrict, fst_conv, snd_conv, split, Pair_iff, singleton_iff,
- converse_iff, domain_converse, range_converse,
- Sigma_empty1, Sigma_empty2, triv_RepFun, subset_refl];
-
-(*Sigma_cong, Pi_cong NOT included by default since they cause
- flex-flex pairs and the "Check your prover" error -- because most
- Sigma's and Pi's are abbreviated as * or -> *)
-val ZF_congs =
- [ball_cong, bex_cong, Replace_cong, RepFun_cong, Collect_cong, lam_cong];
-
val ZF_atomize = atomize (ZF_conn_pairs, ZF_mem_pairs);
-val ZF_ss = FOL_ss
- setmksimps (map mk_meta_eq o ZF_atomize o gen_all)
- addsimps (ZF_simps @ mem_simps @ bquant_simps @
- Collect_simps @ UnInt_simps)
- addcongs ZF_congs;
+simpset := !simpset setmksimps (map mk_meta_eq o ZF_atomize o gen_all);
+
+val ZF_ss = !simpset;
--- a/src/ZF/subset.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/subset.ML Fri Jan 03 15:01:55 1997 +0100
@@ -9,43 +9,38 @@
(*** cons ***)
-qed_goal "cons_subsetI" ZF.thy "[| a:C; B<=C |] ==> cons(a,B) <= C"
- (fn prems=>
- [ (cut_facts_tac prems 1),
- (REPEAT (ares_tac [subsetI] 1
- ORELSE eresolve_tac [consE,ssubst,subsetD] 1)) ]);
+qed_goal "cons_subsetI" thy "!!a. [| a:C; B<=C |] ==> cons(a,B) <= C"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "subset_consI" ZF.thy "B <= cons(a,B)"
- (fn _=> [ (rtac subsetI 1), (etac consI2 1) ]);
+qed_goal "subset_consI" thy "B <= cons(a,B)"
+ (fn _ => [ Fast_tac 1 ]);
-(*Useful for rewriting!*)
-qed_goal "cons_subset_iff" ZF.thy "cons(a,B)<=C <-> a:C & B<=C"
- (fn _=> [ (fast_tac upair_cs 1) ]);
+qed_goal "cons_subset_iff" thy "cons(a,B)<=C <-> a:C & B<=C"
+ (fn _ => [ Fast_tac 1 ]);
(*A safe special case of subset elimination, adding no new variables
[| cons(a,B) <= C; [| a : C; B <= C |] ==> R |] ==> R *)
-bind_thm ("cons_subsetE", (cons_subset_iff RS iffD1 RS conjE));
+bind_thm ("cons_subsetE", cons_subset_iff RS iffD1 RS conjE);
-qed_goal "subset_empty_iff" ZF.thy "A<=0 <-> A=0"
- (fn _=> [ (fast_tac (upair_cs addIs [equalityI]) 1) ]);
+qed_goal "subset_empty_iff" thy "A<=0 <-> A=0"
+ (fn _=> [ (fast_tac (!claset addIs [equalityI]) 1) ]);
-qed_goal "subset_cons_iff" ZF.thy
+qed_goal "subset_cons_iff" thy
"C<=cons(a,B) <-> C<=B | (a:C & C-{a} <= B)"
- (fn _=> [ (fast_tac upair_cs 1) ]);
+ (fn _=> [ (Fast_tac 1) ]);
(*** succ ***)
-qed_goal "subset_succI" ZF.thy "i <= succ(i)"
- (fn _=> [ (rtac subsetI 1), (etac succI2 1) ]);
+qed_goal "subset_succI" thy "i <= succ(i)"
+ (fn _=> [ (Fast_tac 1) ]);
(*But if j is an ordinal or is transitive, then i:j implies i<=j!
See ordinal/Ord_succ_subsetI*)
-qed_goalw "succ_subsetI" ZF.thy [succ_def]
- "[| i:j; i<=j |] ==> succ(i)<=j"
- (fn prems=>
- [ (REPEAT (ares_tac (prems@[cons_subsetI]) 1)) ]);
+qed_goalw "succ_subsetI" thy [succ_def]
+ "!!i j. [| i:j; i<=j |] ==> succ(i)<=j"
+ (fn _=> [ (Fast_tac 1) ]);
-qed_goalw "succ_subsetE" ZF.thy [succ_def]
+qed_goalw "succ_subsetE" thy [succ_def]
"[| succ(i) <= j; [| i:j; i<=j |] ==> P \
\ |] ==> P"
(fn major::prems=>
@@ -54,25 +49,22 @@
(*** singletons ***)
-qed_goal "singleton_subsetI" ZF.thy
- "a:C ==> {a} <= C"
- (fn prems=>
- [ (REPEAT (resolve_tac (prems@[cons_subsetI,empty_subsetI]) 1)) ]);
+qed_goal "singleton_subsetI" thy "!!a c. a:C ==> {a} <= C"
+ (fn _=> [ (Fast_tac 1) ]);
-qed_goal "singleton_subsetD" ZF.thy
- "{a} <= C ==> a:C"
- (fn prems=> [ (REPEAT (ares_tac (prems@[cons_subsetE]) 1)) ]);
+qed_goal "singleton_subsetD" thy "!!a C. {a} <= C ==> a:C"
+ (fn _=> [ (Fast_tac 1) ]);
(*** Big Union -- least upper bound of a set ***)
-qed_goal "Union_subset_iff" ZF.thy "Union(A) <= C <-> (ALL x:A. x <= C)"
- (fn _ => [ fast_tac upair_cs 1 ]);
+qed_goal "Union_subset_iff" thy "Union(A) <= C <-> (ALL x:A. x <= C)"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Union_upper" ZF.thy
- "B:A ==> B <= Union(A)"
- (fn prems=> [ (REPEAT (ares_tac (prems@[subsetI,UnionI]) 1)) ]);
+qed_goal "Union_upper" thy
+ "!!B A. B:A ==> B <= Union(A)"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Union_least" ZF.thy
+qed_goal "Union_least" thy
"[| !!x. x:A ==> x<=C |] ==> Union(A) <= C"
(fn [prem]=>
[ (rtac (ballI RS (Union_subset_iff RS iffD2)) 1),
@@ -80,19 +72,19 @@
(*** Union of a family of sets ***)
-goal ZF.thy "A <= (UN i:I. B(i)) <-> A = (UN i:I. A Int B(i))";
-by (fast_tac (upair_cs addSIs [equalityI] addSEs [equalityE]) 1);
+goal thy "A <= (UN i:I. B(i)) <-> A = (UN i:I. A Int B(i))";
+by (fast_tac (!claset addSIs [equalityI] addSEs [equalityE]) 1);
qed "subset_UN_iff_eq";
-qed_goal "UN_subset_iff" ZF.thy
+qed_goal "UN_subset_iff" thy
"(UN x:A.B(x)) <= C <-> (ALL x:A. B(x) <= C)"
- (fn _ => [ fast_tac upair_cs 1 ]);
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "UN_upper" ZF.thy
+qed_goal "UN_upper" thy
"!!x A. x:A ==> B(x) <= (UN x:A.B(x))"
(fn _ => [ etac (RepFunI RS Union_upper) 1 ]);
-qed_goal "UN_least" ZF.thy
+qed_goal "UN_least" thy
"[| !!x. x:A ==> B(x)<=C |] ==> (UN x:A.B(x)) <= C"
(fn [prem]=>
[ (rtac (ballI RS (UN_subset_iff RS iffD2)) 1),
@@ -101,16 +93,14 @@
(*** Big Intersection -- greatest lower bound of a nonempty set ***)
-qed_goal "Inter_subset_iff" ZF.thy
+qed_goal "Inter_subset_iff" thy
"!!a A. a: A ==> C <= Inter(A) <-> (ALL x:A. C <= x)"
- (fn _ => [ fast_tac upair_cs 1 ]);
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Inter_lower" ZF.thy "B:A ==> Inter(A) <= B"
- (fn prems=>
- [ (REPEAT (resolve_tac (prems@[subsetI]) 1
- ORELSE etac InterD 1)) ]);
+qed_goal "Inter_lower" thy "!!B A. B:A ==> Inter(A) <= B"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Inter_greatest" ZF.thy
+qed_goal "Inter_greatest" thy
"[| a:A; !!x. x:A ==> C<=x |] ==> C <= Inter(A)"
(fn [prem1,prem2]=>
[ (rtac ([prem1, ballI] MRS (Inter_subset_iff RS iffD2)) 1),
@@ -118,12 +108,11 @@
(*** Intersection of a family of sets ***)
-qed_goal "INT_lower" ZF.thy
- "x:A ==> (INT x:A.B(x)) <= B(x)"
- (fn [prem] =>
- [ rtac (prem RS RepFunI RS Inter_lower) 1 ]);
+qed_goal "INT_lower" thy
+ "!!x. x:A ==> (INT x:A.B(x)) <= B(x)"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "INT_greatest" ZF.thy
+qed_goal "INT_greatest" thy
"[| a:A; !!x. x:A ==> C<=B(x) |] ==> C <= (INT x:A.B(x))"
(fn [nonempty,prem] =>
[ rtac (nonempty RS RepFunI RS Inter_greatest) 1,
@@ -132,68 +121,67 @@
(*** Finite Union -- the least upper bound of 2 sets ***)
-qed_goal "Un_subset_iff" ZF.thy "A Un B <= C <-> A <= C & B <= C"
- (fn _ => [ fast_tac upair_cs 1 ]);
+qed_goal "Un_subset_iff" thy "A Un B <= C <-> A <= C & B <= C"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Un_upper1" ZF.thy "A <= A Un B"
- (fn _ => [ (REPEAT (ares_tac [subsetI,UnI1] 1)) ]);
+qed_goal "Un_upper1" thy "A <= A Un B"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Un_upper2" ZF.thy "B <= A Un B"
- (fn _ => [ (REPEAT (ares_tac [subsetI,UnI2] 1)) ]);
+qed_goal "Un_upper2" thy "B <= A Un B"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Un_least" ZF.thy "!!A B C. [| A<=C; B<=C |] ==> A Un B <= C"
- (fn _ =>
- [ (rtac (Un_subset_iff RS iffD2) 1),
- (REPEAT (ares_tac [conjI] 1)) ]);
+qed_goal "Un_least" thy "!!A B C. [| A<=C; B<=C |] ==> A Un B <= C"
+ (fn _ => [ Fast_tac 1 ]);
+
(*** Finite Intersection -- the greatest lower bound of 2 sets *)
-qed_goal "Int_subset_iff" ZF.thy "C <= A Int B <-> C <= A & C <= B"
- (fn _ => [ fast_tac upair_cs 1 ]);
+qed_goal "Int_subset_iff" thy "C <= A Int B <-> C <= A & C <= B"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Int_lower1" ZF.thy "A Int B <= A"
- (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)) ]);
+qed_goal "Int_lower1" thy "A Int B <= A"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Int_lower2" ZF.thy "A Int B <= B"
- (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac IntE 1)) ]);
+qed_goal "Int_lower2" thy "A Int B <= B"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Int_greatest" ZF.thy "!!A B C. [| C<=A; C<=B |] ==> C <= A Int B"
- (fn prems=>
- [ (rtac (Int_subset_iff RS iffD2) 1),
- (REPEAT (ares_tac [conjI] 1)) ]);
+qed_goal "Int_greatest" thy "!!A B C. [| C<=A; C<=B |] ==> C <= A Int B"
+ (fn _ => [ Fast_tac 1 ]);
+
(*** Set difference *)
-qed_goal "Diff_subset" ZF.thy "A-B <= A"
- (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac DiffE 1)) ]);
+qed_goal "Diff_subset" thy "A-B <= A"
+ (fn _ => [ Fast_tac 1 ]);
-qed_goal "Diff_contains" ZF.thy
- "[| C<=A; C Int B = 0 |] ==> C <= A-B"
- (fn prems=>
- [ (cut_facts_tac prems 1),
- (rtac subsetI 1),
- (REPEAT (ares_tac [DiffI,IntI,notI] 1
- ORELSE eresolve_tac [subsetD,equals0D] 1)) ]);
+qed_goal "Diff_contains" thy
+ "!!C. [| C<=A; C Int B = 0 |] ==> C <= A-B"
+ (fn _ => [ (fast_tac (!claset addSEs [equalityE]) 1) ]);
+
(** Collect **)
-qed_goal "Collect_subset" ZF.thy "Collect(A,P) <= A"
- (fn _ => [ (REPEAT (ares_tac [subsetI] 1 ORELSE etac CollectD1 1)) ]);
+qed_goal "Collect_subset" thy "Collect(A,P) <= A"
+ (fn _ => [ Fast_tac 1 ]);
+
(** RepFun **)
-val prems = goal ZF.thy "[| !!x. x:A ==> f(x): B |] ==> {f(x). x:A} <= B";
-by (rtac subsetI 1);
-by (etac RepFunE 1);
-by (etac ssubst 1);
-by (eresolve_tac prems 1);
+val prems = goal thy "[| !!x. x:A ==> f(x): B |] ==> {f(x). x:A} <= B";
+by (fast_tac (!claset addIs prems) 1);
qed "RepFun_subset";
-(*A more powerful claset for subset reasoning*)
-val subset_cs = subset0_cs
- addSIs [subset_refl,cons_subsetI,subset_consI,Union_least,UN_least,Un_least,
- Inter_greatest,Int_greatest,RepFun_subset]
- addSIs [Un_upper1,Un_upper2,Int_lower1,Int_lower2]
- addIs [Union_upper,Inter_lower]
- addSEs [cons_subsetE];
+val subset_SIs =
+ [subset_refl, cons_subsetI, subset_consI,
+ Union_least, UN_least, Un_least,
+ Inter_greatest, Int_greatest, RepFun_subset,
+ Un_upper1, Un_upper2, Int_lower1, Int_lower2];
+
+(*A claset for subset reasoning*)
+val subset_cs = !claset
+ delrules [subsetI, subsetCE]
+ addSIs subset_SIs
+ addIs [Union_upper, Inter_lower]
+ addSEs [cons_subsetE];
+
--- a/src/ZF/subset.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/subset.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,3 @@
(*Dummy theory to document dependencies *)
-subset = "upair"
+subset = upair
--- a/src/ZF/thy_data.ML Fri Jan 03 10:48:28 1997 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,14 +0,0 @@
-(* Title: ZF/thy_data.ML
- ID: $Id$
- Author: Stefan Berghofer
- Copyright 1996 TU Muenchen
-
-Definitions that have to be reread after init_thy_reader has been invoked
-*)
-
-fun claset_of tname =
- case get_thydata tname "claset" of
- None => empty_cs
- | Some (CS_DATA cs) => cs;
-
-
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/ZF/typechk.ML Fri Jan 03 15:01:55 1997 +0100
@@ -0,0 +1,33 @@
+(* Title: ZF/typechk
+ ID: $Id$
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Copyright 1991 University of Cambridge
+
+Tactics for type checking -- from CTT
+*)
+
+fun is_rigid_elem (Const("Trueprop",_) $ (Const("op :",_) $ a $ _)) =
+ not (is_Var (head_of a))
+ | is_rigid_elem _ = false;
+
+(*Try solving a:A by assumption provided a is rigid!*)
+val test_assume_tac = SUBGOAL(fn (prem,i) =>
+ if is_rigid_elem (Logic.strip_assums_concl prem)
+ then assume_tac i else eq_assume_tac i);
+
+(*Type checking solves a:?A (a rigid, ?A maybe flexible).
+ match_tac is too strict; would refuse to instantiate ?A*)
+fun typechk_step_tac tyrls =
+ FIRSTGOAL (test_assume_tac ORELSE' filt_resolve_tac tyrls 3);
+
+fun typechk_tac tyrls = REPEAT (typechk_step_tac tyrls);
+
+val ZF_typechecks =
+ [if_type, lam_type, SigmaI, apply_type, split_type, consI1];
+
+(*Instantiates variables in typing conditions.
+ drawback: does not simplify conjunctions*)
+fun type_auto_tac tyrls hyps = SELECT_GOAL
+ (DEPTH_SOLVE (typechk_step_tac (tyrls@hyps)
+ ORELSE ares_tac [TrueI,refl,iff_refl,ballI,allI,conjI,impI] 1));
+
--- a/src/ZF/upair.ML Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/upair.ML Fri Jan 03 15:01:55 1997 +0100
@@ -21,39 +21,49 @@
(*** Unordered pairs - Upair ***)
-qed_goalw "pairing" ZF.thy [Upair_def]
+qed_goalw "Upair_iff" thy [Upair_def]
"c : Upair(a,b) <-> (c=a | c=b)"
- (fn _ => [ (fast_tac (lemmas_cs addEs [Pow_neq_0, sym RS Pow_neq_0]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addEs [Pow_neq_0, sym RS Pow_neq_0]) 1) ]);
+
+Addsimps [Upair_iff];
-qed_goal "UpairI1" ZF.thy "a : Upair(a,b)"
- (fn _ => [ (rtac (refl RS disjI1 RS (pairing RS iffD2)) 1) ]);
+qed_goal "UpairI1" thy "a : Upair(a,b)"
+ (fn _ => [ Simp_tac 1 ]);
-qed_goal "UpairI2" ZF.thy "b : Upair(a,b)"
- (fn _ => [ (rtac (refl RS disjI2 RS (pairing RS iffD2)) 1) ]);
+qed_goal "UpairI2" thy "b : Upair(a,b)"
+ (fn _ => [ Simp_tac 1 ]);
-qed_goal "UpairE" ZF.thy
+qed_goal "UpairE" thy
"[| a : Upair(b,c); a=b ==> P; a=c ==> P |] ==> P"
(fn major::prems=>
- [ (rtac (major RS (pairing RS iffD1 RS disjE)) 1),
+ [ (rtac (major RS (Upair_iff RS iffD1 RS disjE)) 1),
(REPEAT (eresolve_tac prems 1)) ]);
+(*UpairI1/2 should become UpairCI?*)
+AddSIs [UpairI1,UpairI2];
+AddSEs [UpairE];
+
(*** Rules for binary union -- Un -- defined via Upair ***)
-qed_goalw "UnI1" ZF.thy [Un_def] "c : A ==> c : A Un B"
- (fn [prem]=> [ (rtac (prem RS (UpairI1 RS UnionI)) 1) ]);
+qed_goalw "Un_iff" thy [Un_def] "c : A Un B <-> (c:A | c:B)"
+ (fn _ => [ Fast_tac 1 ]);
+
+Addsimps [Un_iff];
-qed_goalw "UnI2" ZF.thy [Un_def] "c : B ==> c : A Un B"
- (fn [prem]=> [ (rtac (prem RS (UpairI2 RS UnionI)) 1) ]);
+qed_goal "UnI1" thy "!!c. c : A ==> c : A Un B"
+ (fn _ => [ Asm_simp_tac 1 ]);
-qed_goalw "UnE" ZF.thy [Un_def]
+qed_goal "UnI2" thy "!!c. c : B ==> c : A Un B"
+ (fn _ => [ Asm_simp_tac 1 ]);
+
+qed_goal "UnE" thy
"[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"
(fn major::prems=>
- [ (rtac (major RS UnionE) 1),
- (etac UpairE 1),
- (REPEAT (EVERY1 [resolve_tac prems, etac subst, assume_tac])) ]);
+ [ (rtac (major RS (Un_iff RS iffD1 RS disjE)) 1),
+ (REPEAT (eresolve_tac prems 1)) ]);
(*Stronger version of the rule above*)
-qed_goal "UnE'" ZF.thy
+qed_goal "UnE'" thy
"[| c : A Un B; c:A ==> P; [| c:B; c~:A |] ==> P |] ==> P"
(fn major::prems =>
[(rtac (major RS UnE) 1),
@@ -63,81 +73,88 @@
(swap_res_tac prems 1),
(etac notnotD 1)]);
-qed_goal "Un_iff" ZF.thy "c : A Un B <-> (c:A | c:B)"
- (fn _ => [ (fast_tac (lemmas_cs addIs [UnI1,UnI2] addSEs [UnE]) 1) ]);
+(*Classical introduction rule: no commitment to A vs B*)
+qed_goal "UnCI" thy "(c ~: B ==> c : A) ==> c : A Un B"
+ (fn prems=>
+ [ Simp_tac 1, fast_tac (!claset addIs prems) 1 ]);
-(*Classical introduction rule: no commitment to A vs B*)
-qed_goal "UnCI" ZF.thy "(c ~: B ==> c : A) ==> c : A Un B"
- (fn [prem]=>
- [ (rtac (disjCI RS (Un_iff RS iffD2)) 1),
- (etac prem 1) ]);
+AddSIs [UnCI];
+AddSEs [UnE];
(*** Rules for small intersection -- Int -- defined via Upair ***)
-qed_goalw "IntI" ZF.thy [Int_def]
- "[| c : A; c : B |] ==> c : A Int B"
- (fn prems=>
- [ (REPEAT (resolve_tac (prems @ [UpairI1,InterI]) 1
- ORELSE eresolve_tac [UpairE, ssubst] 1)) ]);
+qed_goalw "Int_iff" thy [Int_def] "c : A Int B <-> (c:A & c:B)"
+ (fn _ => [ Fast_tac 1 ]);
+
+Addsimps [Int_iff];
+
+qed_goal "IntI" thy "!!c. [| c : A; c : B |] ==> c : A Int B"
+ (fn _ => [ Asm_simp_tac 1 ]);
-qed_goalw "IntD1" ZF.thy [Int_def] "c : A Int B ==> c : A"
- (fn [major]=>
- [ (rtac (UpairI1 RS (major RS InterD)) 1) ]);
+qed_goal "IntD1" thy "!!c. c : A Int B ==> c : A"
+ (fn _ => [ Asm_full_simp_tac 1 ]);
-qed_goalw "IntD2" ZF.thy [Int_def] "c : A Int B ==> c : B"
- (fn [major]=>
- [ (rtac (UpairI2 RS (major RS InterD)) 1) ]);
+qed_goal "IntD2" thy "!!c. c : A Int B ==> c : B"
+ (fn _ => [ Asm_full_simp_tac 1 ]);
-qed_goal "IntE" ZF.thy
+qed_goal "IntE" thy
"[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"
(fn prems=>
[ (resolve_tac prems 1),
(REPEAT (resolve_tac (prems RL [IntD1,IntD2]) 1)) ]);
-qed_goal "Int_iff" ZF.thy "c : A Int B <-> (c:A & c:B)"
- (fn _ => [ (fast_tac (lemmas_cs addSIs [IntI] addSEs [IntE]) 1) ]);
-
+AddSIs [IntI];
+AddSEs [IntE];
(*** Rules for set difference -- defined via Upair ***)
-qed_goalw "DiffI" ZF.thy [Diff_def]
- "[| c : A; c ~: B |] ==> c : A - B"
- (fn prems=> [ (REPEAT (resolve_tac (prems @ [CollectI]) 1)) ]);
+qed_goalw "Diff_iff" thy [Diff_def] "c : A-B <-> (c:A & c~:B)"
+ (fn _ => [ Fast_tac 1 ]);
+
+Addsimps [Diff_iff];
+
+qed_goal "DiffI" thy "!!c. [| c : A; c ~: B |] ==> c : A - B"
+ (fn _ => [ Asm_simp_tac 1 ]);
-qed_goalw "DiffD1" ZF.thy [Diff_def]
- "c : A - B ==> c : A"
- (fn [major]=> [ (rtac (major RS CollectD1) 1) ]);
+qed_goal "DiffD1" thy "!!c. c : A - B ==> c : A"
+ (fn _ => [ Asm_full_simp_tac 1 ]);
-qed_goalw "DiffD2" ZF.thy [Diff_def]
- "c : A - B ==> c ~: B"
- (fn [major]=> [ (rtac (major RS CollectD2) 1) ]);
+qed_goal "DiffD2" thy "!!c. c : A - B ==> c ~: B"
+ (fn _ => [ Asm_full_simp_tac 1 ]);
-qed_goal "DiffE" ZF.thy
+qed_goal "DiffE" thy
"[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P"
(fn prems=>
[ (resolve_tac prems 1),
(REPEAT (ares_tac (prems RL [DiffD1, DiffD2]) 1)) ]);
-qed_goal "Diff_iff" ZF.thy "c : A-B <-> (c:A & c~:B)"
- (fn _ => [ (fast_tac (lemmas_cs addSIs [DiffI] addSEs [DiffE]) 1) ]);
+AddSIs [DiffI];
+AddSEs [DiffE];
(*** Rules for cons -- defined via Un and Upair ***)
-qed_goalw "consI1" ZF.thy [cons_def] "a : cons(a,B)"
- (fn _ => [ (rtac (UpairI1 RS UnI1) 1) ]);
+qed_goalw "cons_iff" thy [cons_def] "a : cons(b,A) <-> (a=b | a:A)"
+ (fn _ => [ Fast_tac 1 ]);
+
+Addsimps [cons_iff];
-qed_goalw "consI2" ZF.thy [cons_def] "a : B ==> a : cons(b,B)"
- (fn [prem]=> [ (rtac (prem RS UnI2) 1) ]);
+qed_goal "consI1" thy "a : cons(a,B)"
+ (fn _ => [ Simp_tac 1 ]);
+
+Addsimps [consI1];
-qed_goalw "consE" ZF.thy [cons_def]
+qed_goal "consI2" thy "!!B. a : B ==> a : cons(b,B)"
+ (fn _ => [ Asm_simp_tac 1 ]);
+
+qed_goal "consE" thy
"[| a : cons(b,A); a=b ==> P; a:A ==> P |] ==> P"
(fn major::prems=>
- [ (rtac (major RS UnE) 1),
+ [ (rtac (major RS (cons_iff RS iffD1 RS disjE)) 1),
(REPEAT (eresolve_tac (prems @ [UpairE]) 1)) ]);
(*Stronger version of the rule above*)
-qed_goal "consE'" ZF.thy
+qed_goal "consE'" thy
"[| a : cons(b,A); a=b ==> P; [| a:A; a~=b |] ==> P |] ==> P"
(fn major::prems =>
[(rtac (major RS consE) 1),
@@ -147,60 +164,57 @@
(swap_res_tac prems 1),
(etac notnotD 1)]);
-qed_goal "cons_iff" ZF.thy "a : cons(b,A) <-> (a=b | a:A)"
- (fn _ => [ (fast_tac (lemmas_cs addIs [consI1,consI2] addSEs [consE]) 1) ]);
+(*Classical introduction rule*)
+qed_goal "consCI" thy "(a~:B ==> a=b) ==> a: cons(b,B)"
+ (fn prems=>
+ [ Simp_tac 1, fast_tac (!claset addIs prems) 1 ]);
-(*Classical introduction rule*)
-qed_goal "consCI" ZF.thy "(a~:B ==> a=b) ==> a: cons(b,B)"
- (fn [prem]=>
- [ (rtac (disjCI RS (cons_iff RS iffD2)) 1),
- (etac prem 1) ]);
+AddSIs [consCI];
+AddSEs [consE];
-qed_goal "cons_neq_0" ZF.thy "cons(a,B)=0 ==> P"
- (fn [major]=>
- [ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1),
- (rtac consI1 1) ]);
+qed_goal "cons_not_0" thy "cons(a,B) ~= 0"
+ (fn _ => [ (fast_tac (!claset addEs [equalityE]) 1) ]);
-(*Useful for rewriting*)
-qed_goal "cons_not_0" ZF.thy "cons(a,B) ~= 0"
- (fn _=> [ (rtac notI 1), (etac cons_neq_0 1) ]);
+bind_thm ("cons_neq_0", cons_not_0 RS notE);
+
+Addsimps [cons_not_0, cons_not_0 RS not_sym];
+
(*** Singletons - using cons ***)
-qed_goal "singletonI" ZF.thy "a : {a}"
+qed_goal "singleton_iff" thy "a : {b} <-> a=b"
+ (fn _ => [ Simp_tac 1 ]);
+
+qed_goal "singletonI" thy "a : {a}"
(fn _=> [ (rtac consI1 1) ]);
-qed_goal "singletonE" ZF.thy "[| a: {b}; a=b ==> P |] ==> P"
- (fn major::prems=>
- [ (rtac (major RS consE) 1),
- (REPEAT (eresolve_tac (prems @ [emptyE]) 1)) ]);
+bind_thm ("singletonE", make_elim (singleton_iff RS iffD1));
-qed_goal "singleton_iff" ZF.thy "a : {b} <-> a=b"
- (fn _ => [ (fast_tac (lemmas_cs addIs [consI1] addSEs [consE]) 1) ]);
-
+AddSIs [singletonI];
+AddSEs [singletonE];
(*** Rules for Descriptions ***)
-qed_goalw "the_equality" ZF.thy [the_def]
+qed_goalw "the_equality" thy [the_def]
"[| P(a); !!x. P(x) ==> x=a |] ==> (THE x. P(x)) = a"
(fn [pa,eq] =>
- [ (fast_tac (lemmas_cs addSIs [singletonI,pa] addIs [equalityI]
- addEs [eq RS subst]) 1) ]);
+ [ (fast_tac (!claset addSIs [pa] addIs [equalityI]
+ addEs [eq RS subst]) 1) ]);
(* Only use this if you already know EX!x. P(x) *)
-qed_goal "the_equality2" ZF.thy
+qed_goal "the_equality2" thy
"!!P. [| EX! x. P(x); P(a) |] ==> (THE x. P(x)) = a"
(fn _ =>
- [ (deepen_tac (lemmas_cs addSIs [the_equality]) 1 1) ]);
+ [ (deepen_tac (!claset addSIs [the_equality]) 1 1) ]);
-qed_goal "theI" ZF.thy "EX! x. P(x) ==> P(THE x. P(x))"
+qed_goal "theI" thy "EX! x. P(x) ==> P(THE x. P(x))"
(fn [major]=>
[ (rtac (major RS ex1E) 1),
(resolve_tac [major RS the_equality2 RS ssubst] 1),
(REPEAT (assume_tac 1)) ]);
(*Easier to apply than theI: conclusion has only one occurrence of P*)
-qed_goal "theI2" ZF.thy
+qed_goal "theI2" thy
"[| EX! x. P(x); !!x. P(x) ==> Q(x) |] ==> Q(THE x.P(x))"
(fn prems => [ resolve_tac prems 1,
rtac theI 1,
@@ -210,109 +224,113 @@
(THE x.P(x)) rewrites to (THE x. Q(x)) *)
(*If it's "undefined", it's zero!*)
-qed_goalw "the_0" ZF.thy [the_def]
+qed_goalw "the_0" thy [the_def]
"!!P. ~ (EX! x. P(x)) ==> (THE x. P(x))=0"
- (fn _ =>
- [ (fast_tac (lemmas_cs addIs [equalityI] addSEs [ReplaceE]) 1) ]);
+ (fn _ => [ (fast_tac (!claset addIs [equalityI] addSEs [ReplaceE]) 1) ]);
(*** if -- a conditional expression for formulae ***)
-goalw ZF.thy [if_def] "if(True,a,b) = a";
-by (fast_tac (lemmas_cs addIs [the_equality]) 1);
+goalw thy [if_def] "if(True,a,b) = a";
+by (fast_tac (!claset addIs [the_equality]) 1);
qed "if_true";
-goalw ZF.thy [if_def] "if(False,a,b) = b";
-by (fast_tac (lemmas_cs addIs [the_equality]) 1);
+goalw thy [if_def] "if(False,a,b) = b";
+by (fast_tac (!claset addIs [the_equality]) 1);
qed "if_false";
(*Never use with case splitting, or if P is known to be true or false*)
-val prems = goalw ZF.thy [if_def]
+val prems = goalw thy [if_def]
"[| P<->Q; Q ==> a=c; ~Q ==> b=d |] ==> if(P,a,b) = if(Q,c,d)";
-by (simp_tac (FOL_ss addsimps prems addcongs [conj_cong]) 1);
+by (simp_tac (!simpset addsimps prems addcongs [conj_cong]) 1);
qed "if_cong";
(*Not needed for rewriting, since P would rewrite to True anyway*)
-goalw ZF.thy [if_def] "!!P. P ==> if(P,a,b) = a";
-by (fast_tac (lemmas_cs addSIs [the_equality]) 1);
+goalw thy [if_def] "!!P. P ==> if(P,a,b) = a";
+by (fast_tac (!claset addSIs [the_equality]) 1);
qed "if_P";
(*Not needed for rewriting, since P would rewrite to False anyway*)
-goalw ZF.thy [if_def] "!!P. ~P ==> if(P,a,b) = b";
-by (fast_tac (lemmas_cs addSIs [the_equality]) 1);
+goalw thy [if_def] "!!P. ~P ==> if(P,a,b) = b";
+by (fast_tac (!claset addSIs [the_equality]) 1);
qed "if_not_P";
-val if_ss = FOL_ss addsimps [if_true,if_false];
+Addsimps [if_true, if_false];
-qed_goal "expand_if" ZF.thy
+qed_goal "expand_if" thy
"P(if(Q,x,y)) <-> ((Q --> P(x)) & (~Q --> P(y)))"
(fn _=> [ (excluded_middle_tac "Q" 1),
- (asm_simp_tac if_ss 1),
- (asm_simp_tac if_ss 1) ]);
+ (Asm_simp_tac 1),
+ (Asm_simp_tac 1) ]);
-qed_goal "if_iff" ZF.thy "a: if(P,x,y) <-> P & a:x | ~P & a:y"
- (fn _=> [ (simp_tac (if_ss setloop split_tac [expand_if]) 1) ]);
+qed_goal "if_iff" thy "a: if(P,x,y) <-> P & a:x | ~P & a:y"
+ (fn _=> [ (simp_tac (!simpset setloop split_tac [expand_if]) 1) ]);
-qed_goal "if_type" ZF.thy
+qed_goal "if_type" thy
"[| P ==> a: A; ~P ==> b: A |] ==> if(P,a,b): A"
(fn prems=> [ (simp_tac
- (if_ss addsimps prems setloop split_tac [expand_if]) 1) ]);
+ (!simpset addsimps prems setloop split_tac [expand_if]) 1) ]);
(*** Foundation lemmas ***)
(*was called mem_anti_sym*)
-qed_goal "mem_asym" ZF.thy "!!P. [| a:b; b:a |] ==> P"
+qed_goal "mem_asym" thy "!!P. [| a:b; b:a |] ==> P"
(fn _=>
[ (res_inst_tac [("A1","{a,b}")] (foundation RS disjE) 1),
- (etac equals0D 1),
- (rtac consI1 1),
- (fast_tac (lemmas_cs addIs [consI1,consI2]
- addSEs [consE,equalityE]) 1) ]);
+ REPEAT (fast_tac (!claset addSEs [equalityE]) 1) ]);
(*was called mem_anti_refl*)
-qed_goal "mem_irrefl" ZF.thy "a:a ==> P"
- (fn [major]=> [ (rtac (major RS (major RS mem_asym)) 1) ]);
+qed_goal "mem_irrefl" thy "a:a ==> P"
+ (fn [major]=> [ (rtac ([major,major] MRS mem_asym) 1) ]);
-qed_goal "mem_not_refl" ZF.thy "a ~: a"
+(*mem_irrefl should NOT be added to default databases:
+ it would be tried on most goals, making proofs slower!*)
+
+qed_goal "mem_not_refl" thy "a ~: a"
(K [ (rtac notI 1), (etac mem_irrefl 1) ]);
(*Good for proving inequalities by rewriting*)
-qed_goal "mem_imp_not_eq" ZF.thy "!!a A. a:A ==> a ~= A"
- (fn _=> [ fast_tac (lemmas_cs addSEs [mem_irrefl]) 1 ]);
+qed_goal "mem_imp_not_eq" thy "!!a A. a:A ==> a ~= A"
+ (fn _=> [ fast_tac (!claset addSEs [mem_irrefl]) 1 ]);
(*** Rules for succ ***)
-qed_goalw "succI1" ZF.thy [succ_def] "i : succ(i)"
+qed_goalw "succ_iff" thy [succ_def] "i : succ(j) <-> i=j | i:j"
+ (fn _ => [ Fast_tac 1 ]);
+
+qed_goalw "succI1" thy [succ_def] "i : succ(i)"
(fn _=> [ (rtac consI1 1) ]);
-qed_goalw "succI2" ZF.thy [succ_def]
+Addsimps [succI1];
+
+qed_goalw "succI2" thy [succ_def]
"i : j ==> i : succ(j)"
(fn [prem]=> [ (rtac (prem RS consI2) 1) ]);
-qed_goalw "succE" ZF.thy [succ_def]
+qed_goalw "succE" thy [succ_def]
"[| i : succ(j); i=j ==> P; i:j ==> P |] ==> P"
(fn major::prems=>
[ (rtac (major RS consE) 1),
(REPEAT (eresolve_tac prems 1)) ]);
-qed_goal "succ_iff" ZF.thy "i : succ(j) <-> i=j | i:j"
- (fn _ => [ (fast_tac (lemmas_cs addIs [succI1,succI2] addSEs [succE]) 1) ]);
-
(*Classical introduction rule*)
-qed_goal "succCI" ZF.thy "(i~:j ==> i=j) ==> i: succ(j)"
+qed_goal "succCI" thy "(i~:j ==> i=j) ==> i: succ(j)"
(fn [prem]=>
[ (rtac (disjCI RS (succ_iff RS iffD2)) 1),
(etac prem 1) ]);
-qed_goal "succ_neq_0" ZF.thy "succ(n)=0 ==> P"
- (fn [major]=>
- [ (rtac (major RS equalityD1 RS subsetD RS emptyE) 1),
- (rtac succI1 1) ]);
+AddSIs [succCI];
+AddSEs [succE];
+
+qed_goal "succ_not_0" thy "succ(n) ~= 0"
+ (fn _=> [ (fast_tac (!claset addSEs [equalityE]) 1) ]);
-(*Useful for rewriting*)
-qed_goal "succ_not_0" ZF.thy "succ(n) ~= 0"
- (fn _=> [ (rtac notI 1), (etac succ_neq_0 1) ]);
+bind_thm ("succ_neq_0", succ_not_0 RS notE);
+
+Addsimps [succ_not_0, succ_not_0 RS not_sym];
+AddSEs [succ_neq_0, sym RS succ_neq_0];
+
(* succ(c) <= B ==> c : B *)
val succ_subsetD = succI1 RSN (2,subsetD);
@@ -321,18 +339,13 @@
bind_thm ("succ_neq_self", succI1 RS mem_imp_not_eq RS not_sym);
-qed_goal "succ_inject" ZF.thy "!!m n. succ(m) = succ(n) ==> m=n"
- (fn _ =>
- [ (fast_tac (lemmas_cs addSEs [succE, equalityE, make_elim succ_subsetD]
- addEs [mem_asym]) 1) ]);
+qed_goal "succ_inject_iff" thy "succ(m) = succ(n) <-> m=n"
+ (fn _=> [ (fast_tac (!claset addEs [mem_asym, equalityE]) 1) ]);
-qed_goal "succ_inject_iff" ZF.thy "succ(m) = succ(n) <-> m=n"
- (fn _=> [ (fast_tac (FOL_cs addSEs [succ_inject]) 1) ]);
+bind_thm ("succ_inject", succ_inject_iff RS iffD1);
-(*UpairI1/2 should become UpairCI?
- mem_irrefl should NOT be added as an elim-rule here;
- it would be tried on most goals, making proof slower!*)
-val upair_cs = lemmas_cs
- addSIs [singletonI, DiffI, IntI, UnCI, consCI, succCI, UpairI1,UpairI2]
- addSEs [singletonE, DiffE, IntE, UnE, consE, succE, UpairE];
+Addsimps [succ_inject_iff];
+AddSDs [succ_inject];
+
+use"simpdata.ML";
--- a/src/ZF/upair.thy Fri Jan 03 10:48:28 1997 +0100
+++ b/src/ZF/upair.thy Fri Jan 03 15:01:55 1997 +0100
@@ -1,3 +1,21 @@
-(*Dummy theory to document dependencies *)
+(* Title: ZF/upair.thy
+ ID: $Id$
+ Author: Lawrence C Paulson and Martin D Coen, CU Computer Laboratory
+ Copyright 1993 University of Cambridge
+
+Definitions involving unordered pairing
+*)
+
+upair = ZF +
-upair = ZF
+defs
+ (* Definite descriptions -- via Replace over the set "1" *)
+ the_def "The(P) == Union({y . x:{0}, P(y)})"
+ if_def "if(P,a,b) == THE z. P & z=a | ~P & z=b"
+
+ (*Set difference; binary union and intersection*)
+ Diff_def "A - B == { x:A . ~(x:B) }"
+ Un_def "A Un B == Union(Upair(A,B))"
+ Int_def "A Int B == Inter(Upair(A,B))"
+
+end