# HG changeset patch # User wenzelm # Date 1004306352 -3600 # Node ID 0a3dace545c521646a369bf23e37fcd700de71b9 # Parent a394d3e9c8bbd8e415f1e7cd21df8e2618f52e28 converted theory "Set"; diff -r a394d3e9c8bb -r 0a3dace545c5 src/HOL/Set.ML --- a/src/HOL/Set.ML Sun Oct 28 22:58:39 2001 +0100 +++ b/src/HOL/Set.ML Sun Oct 28 22:59:12 2001 +0100 @@ -1,845 +1,26 @@ -(* Title: HOL/Set.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1991 University of Cambridge -Set theory for higher-order logic. A set is simply a predicate. -*) - -section "Relating predicates and sets"; - -Addsimps [Collect_mem_eq]; -AddIffs [mem_Collect_eq]; - -Goal "P(a) ==> a : {x. P(x)}"; -by (Asm_simp_tac 1); -qed "CollectI"; - -Goal "a : {x. P(x)} ==> P(a)"; -by (Asm_full_simp_tac 1); -qed "CollectD"; - -val [prem] = Goal "(!!x. (x:A) = (x:B)) ==> A = B"; -by (rtac (prem RS ext RS arg_cong RS box_equals) 1); -by (rtac Collect_mem_eq 1); -by (rtac Collect_mem_eq 1); -qed "set_ext"; - -val [prem] = Goal "(!!x. P(x)=Q(x)) ==> {x. P(x)} = {x. Q(x)}"; -by (rtac (prem RS ext RS arg_cong) 1); -qed "Collect_cong"; - -bind_thm ("CollectE", make_elim CollectD); - -AddSIs [CollectI]; -AddSEs [CollectE]; - - -section "Bounded quantifiers"; - -val prems = Goalw [Ball_def] - "[| !!x. x:A ==> P(x) |] ==> ALL x:A. P(x)"; -by (REPEAT (ares_tac (prems @ [allI,impI]) 1)); -qed "ballI"; - -bind_thms ("strip", [impI, allI, ballI]); - -Goalw [Ball_def] "[| ALL x:A. P(x); x:A |] ==> P(x)"; -by (Blast_tac 1); -qed "bspec"; - -val major::prems = Goalw [Ball_def] - "[| ALL x:A. P(x); P(x) ==> Q; x~:A ==> Q |] ==> Q"; -by (rtac (major RS spec RS impCE) 1); -by (REPEAT (eresolve_tac prems 1)); -qed "ballE"; - -(*Takes assumptions ALL x:A.P(x) and a:A; creates assumption P(a)*) -fun ball_tac i = etac ballE i THEN contr_tac (i+1); - -AddSIs [ballI]; -AddEs [ballE]; -AddXDs [bspec]; -(* gives better instantiation for bound: *) -claset_ref() := claset() addbefore ("bspec", datac bspec 1); - -(*Normally the best argument order: P(x) constrains the choice of x:A*) -Goalw [Bex_def] "[| P(x); x:A |] ==> EX x:A. P(x)"; -by (Blast_tac 1); -qed "bexI"; - -(*The best argument order when there is only one x:A*) -Goalw [Bex_def] "[| x:A; P(x) |] ==> EX x:A. P(x)"; -by (Blast_tac 1); -qed "rev_bexI"; - -val prems = Goal - "[| ALL x:A. ~P(x) ==> P(a); a:A |] ==> EX x:A. P(x)"; -by (rtac classical 1); -by (REPEAT (ares_tac (prems@[bexI,ballI,notI,notE]) 1)) ; -qed "bexCI"; - -val major::prems = Goalw [Bex_def] - "[| EX x:A. P(x); !!x. [| x:A; P(x) |] ==> Q |] ==> Q"; -by (rtac (major RS exE) 1); -by (REPEAT (eresolve_tac (prems @ [asm_rl,conjE]) 1)); -qed "bexE"; - -AddIs [bexI]; -AddSEs [bexE]; - -(*Trival rewrite rule*) -Goal "(ALL x:A. P) = ((EX x. x:A) --> P)"; -by (simp_tac (simpset() addsimps [Ball_def]) 1); -qed "ball_triv"; - -(*Dual form for existentials*) -Goal "(EX x:A. P) = ((EX x. x:A) & P)"; -by (simp_tac (simpset() addsimps [Bex_def]) 1); -qed "bex_triv"; - -Addsimps [ball_triv, bex_triv]; - -Goal "(EX x:A. x=a) = (a:A)"; -by(Blast_tac 1); -qed "bex_triv_one_point1"; - -Goal "(EX x:A. a=x) = (a:A)"; -by(Blast_tac 1); -qed "bex_triv_one_point2"; - -Goal "(EX x:A. x=a & P x) = (a:A & P a)"; -by(Blast_tac 1); -qed "bex_one_point1"; - -Goal "(EX x:A. a=x & P x) = (a:A & P a)"; -by(Blast_tac 1); -qed "bex_one_point2"; - -Goal "(ALL x:A. x=a --> P x) = (a:A --> P a)"; -by(Blast_tac 1); -qed "ball_one_point1"; - -Goal "(ALL x:A. a=x --> P x) = (a:A --> P a)"; -by(Blast_tac 1); -qed "ball_one_point2"; - -Addsimps [bex_triv_one_point1,bex_triv_one_point2, - bex_one_point1,bex_one_point2, - ball_one_point1,ball_one_point2]; - -let -val ex_pattern = Thm.read_cterm (Theory.sign_of (the_context ())) - ("EX x:A. P(x) & Q(x)",HOLogic.boolT) - -val prove_bex_tac = rewrite_goals_tac [Bex_def] THEN - Quantifier1.prove_one_point_ex_tac; - -val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac; - -val all_pattern = Thm.read_cterm (Theory.sign_of (the_context ())) - ("ALL x:A. P(x) --> Q(x)",HOLogic.boolT) - -val prove_ball_tac = rewrite_goals_tac [Ball_def] THEN - Quantifier1.prove_one_point_all_tac; - -val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac; - -val defBEX_regroup = mk_simproc "defined BEX" [ex_pattern] rearrange_bex; -val defBALL_regroup = mk_simproc "defined BALL" [all_pattern] rearrange_ball; -in - -Addsimprocs [defBALL_regroup,defBEX_regroup] - -end; - -(** Congruence rules **) - -val prems = Goalw [Ball_def] - "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ -\ (ALL x:A. P(x)) = (ALL x:B. Q(x))"; -by (asm_simp_tac (simpset() addsimps prems) 1); -qed "ball_cong"; - -val prems = Goalw [Bex_def] - "[| A=B; !!x. x:B ==> P(x) = Q(x) |] ==> \ -\ (EX x:A. P(x)) = (EX x:B. Q(x))"; -by (asm_simp_tac (simpset() addcongs [conj_cong] addsimps prems) 1); -qed "bex_cong"; - -Addcongs [ball_cong,bex_cong]; - -section "Subsets"; - -val prems = Goalw [subset_def] "(!!x. x:A ==> x:B) ==> A <= B"; -by (REPEAT (ares_tac (prems @ [ballI]) 1)); -qed "subsetI"; - -(*Map the type ('a set => anything) to just 'a. - For overloading constants whose first argument has type "'a set" *) -fun overload_1st_set s = Blast.overloaded (s, HOLogic.dest_setT o domain_type); - -(*While (:) is not, its type must be kept - for overloading of = to work.*) -Blast.overloaded ("op :", domain_type); - -overload_1st_set "Ball"; (*need UNION, INTER also?*) -overload_1st_set "Bex"; - -(*Image: retain the type of the set being expressed*) -Blast.overloaded ("image", domain_type); - -(*Rule in Modus Ponens style*) -Goalw [subset_def] "[| A <= B; c:A |] ==> c:B"; -by (Blast_tac 1); -qed "subsetD"; -AddXIs [subsetD]; - -(*The same, with reversed premises for use with etac -- cf rev_mp*) -Goal "[| c:A; A <= B |] ==> c:B"; -by (REPEAT (ares_tac [subsetD] 1)) ; -qed "rev_subsetD"; -AddXIs [rev_subsetD]; - -(*Converts A<=B to x:A ==> x:B*) -fun impOfSubs th = th RSN (2, rev_subsetD); - -(*Classical elimination rule*) -val major::prems = Goalw [subset_def] - "[| A <= B; c~:A ==> P; c:B ==> P |] ==> P"; -by (rtac (major RS ballE) 1); -by (REPEAT (eresolve_tac prems 1)); -qed "subsetCE"; - -(*Takes assumptions A<=B; c:A and creates the assumption c:B *) -fun set_mp_tac i = etac subsetCE i THEN mp_tac i; - -AddSIs [subsetI]; -AddEs [subsetD, subsetCE]; - -Goal "[| A <= B; c ~: B |] ==> c ~: A"; -by (Blast_tac 1); -qed "contra_subsetD"; - -Goal "A <= (A::'a set)"; -by (Fast_tac 1); -qed "subset_refl"; - -Goal "[| A<=B; B<=C |] ==> A<=(C::'a set)"; -by (Blast_tac 1); -qed "subset_trans"; - - -section "Equality"; - -(*Anti-symmetry of the subset relation*) -Goal "[| A <= B; B <= A |] ==> A = (B::'a set)"; -by (rtac set_ext 1); -by (blast_tac (claset() addIs [subsetD]) 1); -qed "subset_antisym"; -bind_thm ("equalityI", subset_antisym); - -AddSIs [equalityI]; - -(* Equality rules from ZF set theory -- are they appropriate here? *) -Goal "A = B ==> A<=(B::'a set)"; -by (etac ssubst 1); -by (rtac subset_refl 1); -qed "equalityD1"; - -Goal "A = B ==> B<=(A::'a set)"; -by (etac ssubst 1); -by (rtac subset_refl 1); -qed "equalityD2"; - -(*Be careful when adding this to the claset as subset_empty is in the simpset: - A={} goes to {}<=A and A<={} and then back to A={} !*) -val prems = Goal - "[| A = B; [| A<=B; B<=(A::'a set) |] ==> P |] ==> P"; -by (resolve_tac prems 1); -by (REPEAT (resolve_tac (prems RL [equalityD1,equalityD2]) 1)); -qed "equalityE"; - -val major::prems = Goal - "[| A = B; [| c:A; c:B |] ==> P; [| c~:A; c~:B |] ==> P |] ==> P"; -by (rtac (major RS equalityE) 1); -by (REPEAT (contr_tac 1 ORELSE eresolve_tac ([asm_rl,subsetCE]@prems) 1)); -qed "equalityCE"; - -AddEs [equalityCE]; - -(*Lemma for creating induction formulae -- for "pattern matching" on p - To make the induction hypotheses usable, apply "spec" or "bspec" to - put universal quantifiers over the free variables in p. *) -val prems = Goal - "[| p:A; !!z. z:A ==> p=z --> R |] ==> R"; -by (rtac mp 1); -by (REPEAT (resolve_tac (refl::prems) 1)); -qed "setup_induction"; - -Goal "A = B ==> (x : A) = (x : B)"; -by (Asm_simp_tac 1); -qed "eqset_imp_iff"; - - -section "The universal set -- UNIV"; - -Goalw [UNIV_def] "x : UNIV"; -by (rtac CollectI 1); -by (rtac TrueI 1); -qed "UNIV_I"; - -Addsimps [UNIV_I]; -AddIs [UNIV_I]; (*unsafe makes it less likely to cause problems*) - -Goal "EX x. x : UNIV"; -by (Simp_tac 1); -qed "UNIV_witness"; -AddXIs [UNIV_witness]; - -Goal "A <= UNIV"; -by (rtac subsetI 1); -by (rtac UNIV_I 1); -qed "subset_UNIV"; - -(** Eta-contracting these two rules (to remove P) causes them to be ignored - because of their interaction with congruence rules. **) - -Goalw [Ball_def] "Ball UNIV P = All P"; -by (Simp_tac 1); -qed "ball_UNIV"; - -Goalw [Bex_def] "Bex UNIV P = Ex P"; -by (Simp_tac 1); -qed "bex_UNIV"; -Addsimps [ball_UNIV, bex_UNIV]; - - -section "The empty set -- {}"; - -Goalw [empty_def] "(c : {}) = False"; -by (Blast_tac 1) ; -qed "empty_iff"; - -Addsimps [empty_iff]; - -Goal "a:{} ==> P"; -by (Full_simp_tac 1); -qed "emptyE"; - -AddSEs [emptyE]; - -Goal "{} <= A"; -by (Blast_tac 1) ; -qed "empty_subsetI"; - -(*One effect is to delete the ASSUMPTION {} <= A*) -AddIffs [empty_subsetI]; - -val [prem]= Goal "[| !!y. y:A ==> False |] ==> A={}"; -by (blast_tac (claset() addIs [prem RS FalseE]) 1) ; -qed "equals0I"; - -(*Use for reasoning about disjointness: A Int B = {} *) -Goal "A={} ==> a ~: A"; -by (Blast_tac 1) ; -qed "equals0D"; - -Goalw [Ball_def] "Ball {} P = True"; -by (Simp_tac 1); -qed "ball_empty"; - -Goalw [Bex_def] "Bex {} P = False"; -by (Simp_tac 1); -qed "bex_empty"; -Addsimps [ball_empty, bex_empty]; - -Goal "UNIV ~= {}"; -by (blast_tac (claset() addEs [equalityE]) 1); -qed "UNIV_not_empty"; -AddIffs [UNIV_not_empty]; - - - -section "The Powerset operator -- Pow"; - -Goalw [Pow_def] "(A : Pow(B)) = (A <= B)"; -by (Asm_simp_tac 1); -qed "Pow_iff"; - -AddIffs [Pow_iff]; - -Goalw [Pow_def] "A <= B ==> A : Pow(B)"; -by (etac CollectI 1); -qed "PowI"; - -Goalw [Pow_def] "A : Pow(B) ==> A<=B"; -by (etac CollectD 1); -qed "PowD"; - - -bind_thm ("Pow_bottom", empty_subsetI RS PowI); (* {}: Pow(B) *) -bind_thm ("Pow_top", subset_refl RS PowI); (* A : Pow(A) *) - - -section "Set complement"; - -Goalw [Compl_def] "(c : -A) = (c~:A)"; -by (Blast_tac 1); -qed "Compl_iff"; - -Addsimps [Compl_iff]; - -val prems = Goalw [Compl_def] "[| c:A ==> False |] ==> c : -A"; -by (REPEAT (ares_tac (prems @ [CollectI,notI]) 1)); -qed "ComplI"; - -(*This form, with negated conclusion, works well with the Classical prover. - Negated assumptions behave like formulae on the right side of the notional - turnstile...*) -Goalw [Compl_def] "c : -A ==> c~:A"; -by (etac CollectD 1); -qed "ComplD"; - -bind_thm ("ComplE", make_elim ComplD); - -AddSIs [ComplI]; -AddSEs [ComplE]; - - -section "Binary union -- Un"; - -Goalw [Un_def] "(c : A Un B) = (c:A | c:B)"; -by (Blast_tac 1); -qed "Un_iff"; -Addsimps [Un_iff]; - -Goal "c:A ==> c : A Un B"; -by (Asm_simp_tac 1); -qed "UnI1"; +(* legacy ML bindings *) -Goal "c:B ==> c : A Un B"; -by (Asm_simp_tac 1); -qed "UnI2"; - -AddXEs [UnI1, UnI2]; - - -(*Classical introduction rule: no commitment to A vs B*) - -val prems = Goal "(c~:B ==> c:A) ==> c : A Un B"; -by (Simp_tac 1); -by (REPEAT (ares_tac (prems@[disjCI]) 1)) ; -qed "UnCI"; - -val major::prems = Goalw [Un_def] - "[| c : A Un B; c:A ==> P; c:B ==> P |] ==> P"; -by (rtac (major RS CollectD RS disjE) 1); -by (REPEAT (eresolve_tac prems 1)); -qed "UnE"; - -AddSIs [UnCI]; -AddSEs [UnE]; - - -section "Binary intersection -- Int"; - -Goalw [Int_def] "(c : A Int B) = (c:A & c:B)"; -by (Blast_tac 1); -qed "Int_iff"; -Addsimps [Int_iff]; - -Goal "[| c:A; c:B |] ==> c : A Int B"; -by (Asm_simp_tac 1); -qed "IntI"; - -Goal "c : A Int B ==> c:A"; -by (Asm_full_simp_tac 1); -qed "IntD1"; - -Goal "c : A Int B ==> c:B"; -by (Asm_full_simp_tac 1); -qed "IntD2"; - -val [major,minor] = Goal - "[| c : A Int B; [| c:A; c:B |] ==> P |] ==> P"; -by (rtac minor 1); -by (rtac (major RS IntD1) 1); -by (rtac (major RS IntD2) 1); -qed "IntE"; - -AddSIs [IntI]; -AddSEs [IntE]; - -section "Set difference"; - -Goalw [set_diff_def] "(c : A-B) = (c:A & c~:B)"; -by (Blast_tac 1); -qed "Diff_iff"; -Addsimps [Diff_iff]; - -Goal "[| c : A; c ~: B |] ==> c : A - B"; -by (Asm_simp_tac 1) ; -qed "DiffI"; - -Goal "c : A - B ==> c : A"; -by (Asm_full_simp_tac 1) ; -qed "DiffD1"; - -Goal "[| c : A - B; c : B |] ==> P"; -by (Asm_full_simp_tac 1) ; -qed "DiffD2"; - -val prems = Goal "[| c : A - B; [| c:A; c~:B |] ==> P |] ==> P"; -by (resolve_tac prems 1); -by (REPEAT (ares_tac (prems RL [DiffD1, DiffD2 RS notI]) 1)) ; -qed "DiffE"; - -AddSIs [DiffI]; -AddSEs [DiffE]; - - -section "Augmenting a set -- insert"; - -Goalw [insert_def] "(a : insert b A) = (a=b | a:A)"; -by (Blast_tac 1); -qed "insert_iff"; -Addsimps [insert_iff]; - -Goal "a : insert a B"; -by (Simp_tac 1); -qed "insertI1"; - -Goal "!!a. a : B ==> a : insert b B"; -by (Asm_simp_tac 1); -qed "insertI2"; - -val major::prems = Goalw [insert_def] - "[| a : insert b A; a=b ==> P; a:A ==> P |] ==> P"; -by (rtac (major RS UnE) 1); -by (REPEAT (eresolve_tac (prems @ [CollectE]) 1)); -qed "insertE"; - -(*Classical introduction rule*) -val prems = Goal "(a~:B ==> a=b) ==> a: insert b B"; -by (Simp_tac 1); -by (REPEAT (ares_tac (prems@[disjCI]) 1)) ; -qed "insertCI"; - -AddSIs [insertCI]; -AddSEs [insertE]; - -Goal "(A <= insert x B) = (if x:A then A-{x} <= B else A<=B)"; -by Auto_tac; -qed "subset_insert_iff"; - -section "Singletons, using insert"; - -Goal "a : {a}"; -by (rtac insertI1 1) ; -qed "singletonI"; - -Goal "b : {a} ==> b=a"; -by (Blast_tac 1); -qed "singletonD"; - -bind_thm ("singletonE", make_elim singletonD); - -Goal "(b : {a}) = (b=a)"; -by (Blast_tac 1); -qed "singleton_iff"; - -Goal "{a}={b} ==> a=b"; -by (blast_tac (claset() addEs [equalityE]) 1); -qed "singleton_inject"; - -(*Redundant? But unlike insertCI, it proves the subgoal immediately!*) -AddSIs [singletonI]; -AddSDs [singleton_inject]; -AddSEs [singletonE]; - -Goal "({b} = insert a A) = (a = b & A <= {b})"; -by (blast_tac (claset() addSEs [equalityE]) 1); -qed "singleton_insert_inj_eq"; - -Goal "(insert a A = {b}) = (a = b & A <= {b})"; -by (blast_tac (claset() addSEs [equalityE]) 1); -qed "singleton_insert_inj_eq'"; - -AddIffs [singleton_insert_inj_eq, singleton_insert_inj_eq']; - -Goal "A <= {x} ==> A={} | A = {x}"; -by (Fast_tac 1); -qed "subset_singletonD"; - -Goal "{x. x=a} = {a}"; -by (Blast_tac 1); -qed "singleton_conv"; -Addsimps [singleton_conv]; - -Goal "{x. a=x} = {a}"; -by (Blast_tac 1); -qed "singleton_conv2"; -Addsimps [singleton_conv2]; - -Goal "[| A - {x} <= B; x : A |] ==> A <= insert x B"; -by(Blast_tac 1); -qed "diff_single_insert"; - - -section "Unions of families -- UNION x:A. B(x) is Union(B`A)"; - -Goalw [UNION_def] "(b: (UN x:A. B(x))) = (EX x:A. b: B(x))"; -by (Blast_tac 1); -qed "UN_iff"; - -Addsimps [UN_iff]; - -(*The order of the premises presupposes that A is rigid; b may be flexible*) -Goal "[| a:A; b: B(a) |] ==> b: (UN x:A. B(x))"; -by Auto_tac; -qed "UN_I"; - -val major::prems = Goalw [UNION_def] - "[| b : (UN x:A. B(x)); !!x.[| x:A; b: B(x) |] ==> R |] ==> R"; -by (rtac (major RS CollectD RS bexE) 1); -by (REPEAT (ares_tac prems 1)); -qed "UN_E"; - -AddIs [UN_I]; -AddSEs [UN_E]; - -val prems = Goalw [UNION_def] - "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ -\ (UN x:A. C(x)) = (UN x:B. D(x))"; -by (asm_simp_tac (simpset() addsimps prems) 1); -qed "UN_cong"; -Addcongs [UN_cong]; - - -section "Intersections of families -- INTER x:A. B(x) is Inter(B`A)"; - -Goalw [INTER_def] "(b: (INT x:A. B(x))) = (ALL x:A. b: B(x))"; -by Auto_tac; -qed "INT_iff"; - -Addsimps [INT_iff]; - -val prems = Goalw [INTER_def] - "(!!x. x:A ==> b: B(x)) ==> b : (INT x:A. B(x))"; -by (REPEAT (ares_tac ([CollectI,ballI] @ prems) 1)); -qed "INT_I"; - -Goal "[| b : (INT x:A. B(x)); a:A |] ==> b: B(a)"; -by Auto_tac; -qed "INT_D"; - -(*"Classical" elimination -- by the Excluded Middle on a:A *) -val major::prems = Goalw [INTER_def] - "[| b : (INT x:A. B(x)); b: B(a) ==> R; a~:A ==> R |] ==> R"; -by (rtac (major RS CollectD RS ballE) 1); -by (REPEAT (eresolve_tac prems 1)); -qed "INT_E"; - -AddSIs [INT_I]; -AddEs [INT_D, INT_E]; - -val prems = Goalw [INTER_def] - "[| A=B; !!x. x:B ==> C(x) = D(x) |] ==> \ -\ (INT x:A. C(x)) = (INT x:B. D(x))"; -by (asm_simp_tac (simpset() addsimps prems) 1); -qed "INT_cong"; -Addcongs [INT_cong]; - - -section "Union"; - -Goalw [Union_def] "(A : Union(C)) = (EX X:C. A:X)"; -by (Blast_tac 1); -qed "Union_iff"; - -Addsimps [Union_iff]; - -(*The order of the premises presupposes that C is rigid; A may be flexible*) -Goal "[| X:C; A:X |] ==> A : Union(C)"; -by Auto_tac; -qed "UnionI"; - -val major::prems = Goalw [Union_def] - "[| A : Union(C); !!X.[| A:X; X:C |] ==> R |] ==> R"; -by (rtac (major RS UN_E) 1); -by (REPEAT (ares_tac prems 1)); -qed "UnionE"; - -AddIs [UnionI]; -AddSEs [UnionE]; - - -section "Inter"; - -Goalw [Inter_def] "(A : Inter(C)) = (ALL X:C. A:X)"; -by (Blast_tac 1); -qed "Inter_iff"; - -Addsimps [Inter_iff]; - -val prems = Goalw [Inter_def] - "[| !!X. X:C ==> A:X |] ==> A : Inter(C)"; -by (REPEAT (ares_tac ([INT_I] @ prems) 1)); -qed "InterI"; - -(*A "destruct" rule -- every X in C contains A as an element, but - A:X can hold when X:C does not! This rule is analogous to "spec". *) -Goal "[| A : Inter(C); X:C |] ==> A:X"; -by Auto_tac; -qed "InterD"; - -(*"Classical" elimination rule -- does not require proving X:C *) -val major::prems = Goalw [Inter_def] - "[| A : Inter(C); X~:C ==> R; A:X ==> R |] ==> R"; -by (rtac (major RS INT_E) 1); -by (REPEAT (eresolve_tac prems 1)); -qed "InterE"; - -AddSIs [InterI]; -AddEs [InterD, InterE]; - - -(*** Image of a set under a function ***) - -(*Frequently b does not have the syntactic form of f(x).*) -Goalw [image_def] "[| b=f(x); x:A |] ==> b : f`A"; -by (Blast_tac 1); -qed "image_eqI"; -Addsimps [image_eqI]; - -bind_thm ("imageI", refl RS image_eqI); - -(*This version's more effective when we already have the required x*) -Goalw [image_def] "[| x:A; b=f(x) |] ==> b : f`A"; -by (Blast_tac 1); -qed "rev_image_eqI"; - -(*The eta-expansion gives variable-name preservation.*) -val major::prems = Goalw [image_def] - "[| b : (%x. f(x))`A; !!x.[| b=f(x); x:A |] ==> P |] ==> P"; -by (rtac (major RS CollectD RS bexE) 1); -by (REPEAT (ares_tac prems 1)); -qed "imageE"; - -AddIs [image_eqI]; -AddSEs [imageE]; - -Goal "f`(A Un B) = f`A Un f`B"; -by (Blast_tac 1); -qed "image_Un"; - -Goal "(z : f`A) = (EX x:A. z = f x)"; -by (Blast_tac 1); -qed "image_iff"; - -(*This rewrite rule would confuse users if made default.*) -Goal "(f`A <= B) = (ALL x:A. f(x): B)"; -by (Blast_tac 1); -qed "image_subset_iff"; - -Goal "(B <= f ` A) = (? AA. AA <= A & B = f ` AA)"; -by Safe_tac; -by (Fast_tac 2); -by (res_inst_tac [("x","{a. a : A & f a : B}")] exI 1); -by (Fast_tac 1); -qed "subset_image_iff"; - -(*Replaces the three steps subsetI, imageE, hyp_subst_tac, but breaks too - many existing proofs.*) -val prems = Goal "(!!x. x:A ==> f(x) : B) ==> f`A <= B"; -by (blast_tac (claset() addIs prems) 1); -qed "image_subsetI"; - -(*** Range of a function -- just a translation for image! ***) - -Goal "b=f(x) ==> b : range(f)"; -by (EVERY1 [etac image_eqI, rtac UNIV_I]); -bind_thm ("range_eqI", UNIV_I RSN (2,image_eqI)); - -bind_thm ("rangeI", UNIV_I RS imageI); - -val [major,minor] = Goal - "[| b : range(%x. f(x)); !!x. b=f(x) ==> P |] ==> P"; -by (rtac (major RS imageE) 1); -by (etac minor 1); -qed "rangeE"; -AddXEs [rangeE]; - - -(*** Set reasoning tools ***) - - -(** Rewrite rules for boolean case-splitting: faster than - addsplits[split_if] -**) - -bind_thm ("split_if_eq1", read_instantiate [("P", "%x. x = ?b")] split_if); -bind_thm ("split_if_eq2", read_instantiate [("P", "%x. ?a = x")] split_if); - -(*Split ifs on either side of the membership relation. - Not for Addsimps -- can cause goals to blow up!*) -bind_thm ("split_if_mem1", inst "P" "%x. x : ?b" split_if); -bind_thm ("split_if_mem2", inst "P" "%x. ?a : x" split_if); - -bind_thms ("split_ifs", [if_bool_eq_conj, split_if_eq1, split_if_eq2, - split_if_mem1, split_if_mem2]); - - -(*Each of these has ALREADY been added to simpset() above.*) -bind_thms ("mem_simps", [insert_iff, empty_iff, Un_iff, Int_iff, Compl_iff, Diff_iff, - mem_Collect_eq, UN_iff, Union_iff, INT_iff, Inter_iff]); - -(*Would like to add these, but the existing code only searches for the - outer-level constant, which in this case is just "op :"; we instead need - to use term-nets to associate patterns with rules. Also, if a rule fails to - apply, then the formula should be kept. - [("uminus", Compl_iff RS iffD1), ("op -", [Diff_iff RS iffD1]), - ("op Int", [IntD1,IntD2]), - ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] - *) -val mksimps_pairs = - [("Ball",[bspec])] @ mksimps_pairs; - -simpset_ref() := simpset() setmksimps (mksimps mksimps_pairs); - -Addsimps[subset_UNIV, subset_refl]; - - -(*** The 'proper subset' relation (<) ***) - -Goalw [psubset_def] "!!A::'a set. [| A <= B; A ~= B |] ==> A A < C"; -by (auto_tac (claset(), simpset() addsimps [psubset_eq])); -qed "psubset_subset_trans"; - -Goal"[| (A::'a set) <= B; B < C|] ==> A < C"; -by (auto_tac (claset(), simpset() addsimps [psubset_eq])); -qed "subset_psubset_trans"; - -Goalw [psubset_def] "A < B ==> EX b. b : (B - A)"; -by (Blast_tac 1); -qed "psubset_imp_ex_mem"; - -Goal "(!!x. x:A ==> P x) == Trueprop (ALL x:A. P x)"; -by (simp_tac (simpset () addsimps [Ball_def, thm "atomize_all", thm "atomize_imp"]) 1); -qed "atomize_ball"; +structure Set = +struct + val thy = the_context (); + val Ball_def = Ball_def; + val Bex_def = Bex_def; + val Collect_mem_eq = Collect_mem_eq; + val Compl_def = Compl_def; + val INTER_def = INTER_def; + val Int_def = Int_def; + val Inter_def = Inter_def; + val Pow_def = Pow_def; + val UNION_def = UNION_def; + val UNIV_def = UNIV_def; + val Un_def = Un_def; + val Union_def = Union_def; + val empty_def = empty_def; + val image_def = image_def; + val insert_def = insert_def; + val mem_Collect_eq = mem_Collect_eq; + val psubset_def = psubset_def; + val set_diff_def = set_diff_def; + val subset_def = subset_def; +end; diff -r a394d3e9c8bb -r 0a3dace545c5 src/HOL/Set.thy --- a/src/HOL/Set.thy Sun Oct 28 22:58:39 2001 +0100 +++ b/src/HOL/Set.thy Sun Oct 28 22:59:12 2001 +0100 @@ -4,70 +4,71 @@ Copyright 1993 University of Cambridge *) -Set = HOL + +header {* Set theory for higher-order logic *} + +theory Set = HOL +files ("subset.ML") ("equalities.ML") ("mono.ML"): + +text {* A set in HOL is simply a predicate. *} -(** Core syntax **) +subsection {* Basic syntax *} global -types - 'a set - -arities - set :: (term) term - -instance - set :: (term) {ord, minus} - -syntax - "op :" :: ['a, 'a set] => bool ("op :") +typedecl 'a set +arities set :: ("term") "term" consts - "{}" :: 'a set ("{}") - UNIV :: 'a set - insert :: ['a, 'a set] => 'a set - Collect :: ('a => bool) => 'a set (*comprehension*) - Int :: ['a set, 'a set] => 'a set (infixl 70) - Un :: ['a set, 'a set] => 'a set (infixl 65) - UNION, INTER :: ['a set, 'a => 'b set] => 'b set (*general*) - Union, Inter :: (('a set) set) => 'a set (*of a set*) - Pow :: 'a set => 'a set set (*powerset*) - Ball, Bex :: ['a set, 'a => bool] => bool (*bounded quantifiers*) - "image" :: ['a => 'b, 'a set] => ('b set) (infixr "`" 90) - (*membership*) - "op :" :: ['a, 'a set] => bool ("(_/ : _)" [50, 51] 50) + "{}" :: "'a set" ("{}") + UNIV :: "'a set" + insert :: "'a => 'a set => 'a set" + Collect :: "('a => bool) => 'a set" -- "comprehension" + Int :: "'a set => 'a set => 'a set" (infixl 70) + Un :: "'a set => 'a set => 'a set" (infixl 65) + UNION :: "'a set => ('a => 'b set) => 'b set" -- "general union" + INTER :: "'a set => ('a => 'b set) => 'b set" -- "general intersection" + Union :: "'a set set => 'a set" -- "union of a set" + Inter :: "'a set set => 'a set" -- "intersection of a set" + Pow :: "'a set => 'a set set" -- "powerset" + Ball :: "'a set => ('a => bool) => bool" -- "bounded universal quantifiers" + Bex :: "'a set => ('a => bool) => bool" -- "bounded existential quantifiers" + image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) + +syntax + "op :" :: "'a => 'a set => bool" ("op :") +consts + "op :" :: "'a => 'a set => bool" ("(_/ : _)" [50, 51] 50) -- "membership" + +local + +instance set :: ("term") ord .. +instance set :: ("term") minus .. -(** Additional concrete syntax **) +subsection {* Additional concrete syntax *} syntax - range :: ('a => 'b) => 'b set (*of function*) - - (* Infix syntax for non-membership *) + range :: "('a => 'b) => 'b set" -- "of function" - "op ~:" :: ['a, 'a set] => bool ("op ~:") - "op ~:" :: ['a, 'a set] => bool ("(_/ ~: _)" [50, 51] 50) - + "op ~:" :: "'a => 'a set => bool" ("op ~:") -- "non-membership" + "op ~:" :: "'a => 'a set => bool" ("(_/ ~: _)" [50, 51] 50) - "@Finset" :: args => 'a set ("{(_)}") - "@Coll" :: [pttrn, bool] => 'a set ("(1{_./ _})") - "@SetCompr" :: ['a, idts, bool] => 'a set ("(1{_ |/_./ _})") - - (* Big Intersection / Union *) + "@Finset" :: "args => 'a set" ("{(_)}") + "@Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") + "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") - "@INTER1" :: [pttrns, 'b set] => 'b set ("(3INT _./ _)" 10) - "@UNION1" :: [pttrns, 'b set] => 'b set ("(3UN _./ _)" 10) - "@INTER" :: [pttrn, 'a set, 'b set] => 'b set ("(3INT _:_./ _)" 10) - "@UNION" :: [pttrn, 'a set, 'b set] => 'b set ("(3UN _:_./ _)" 10) + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" 10) + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" 10) - (* Bounded Quantifiers *) - "_Ball" :: [pttrn, 'a set, bool] => bool ("(3ALL _:_./ _)" [0, 0, 10] 10) - "_Bex" :: [pttrn, 'a set, bool] => bool ("(3EX _:_./ _)" [0, 0, 10] 10) + "_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10) syntax (HOL) - "_Ball" :: [pttrn, 'a set, bool] => bool ("(3! _:_./ _)" [0, 0, 10] 10) - "_Bex" :: [pttrn, 'a set, bool] => bool ("(3? _:_./ _)" [0, 0, 10] 10) + "_Ball" :: "pttrn => 'a set => bool => bool" ("(3! _:_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10) translations "range f" == "f`UNIV" @@ -85,120 +86,811 @@ "EX x:A. P" == "Bex A (%x. P)" syntax ("" output) - "_setle" :: ['a set, 'a set] => bool ("op <=") - "_setle" :: ['a set, 'a set] => bool ("(_/ <= _)" [50, 51] 50) - "_setless" :: ['a set, 'a set] => bool ("op <") - "_setless" :: ['a set, 'a set] => bool ("(_/ < _)" [50, 51] 50) + "_setle" :: "'a set => 'a set => bool" ("op <=") + "_setle" :: "'a set => 'a set => bool" ("(_/ <= _)" [50, 51] 50) + "_setless" :: "'a set => 'a set => bool" ("op <") + "_setless" :: "'a set => 'a set => bool" ("(_/ < _)" [50, 51] 50) syntax (symbols) - "_setle" :: ['a set, 'a set] => bool ("op \\") - "_setle" :: ['a set, 'a set] => bool ("(_/ \\ _)" [50, 51] 50) - "_setless" :: ['a set, 'a set] => bool ("op \\") - "_setless" :: ['a set, 'a set] => bool ("(_/ \\ _)" [50, 51] 50) - "op Int" :: ['a set, 'a set] => 'a set (infixl "\\" 70) - "op Un" :: ['a set, 'a set] => 'a set (infixl "\\" 65) - "op :" :: ['a, 'a set] => bool ("op \\") - "op :" :: ['a, 'a set] => bool ("(_/ \\ _)" [50, 51] 50) - "op ~:" :: ['a, 'a set] => bool ("op \\") - "op ~:" :: ['a, 'a set] => bool ("(_/ \\ _)" [50, 51] 50) - "@UNION1" :: [pttrns, 'b set] => 'b set ("(3\\_./ _)" 10) - "@INTER1" :: [pttrns, 'b set] => 'b set ("(3\\_./ _)" 10) - "@UNION" :: [pttrn, 'a set, 'b set] => 'b set ("(3\\_\\_./ _)" 10) - "@INTER" :: [pttrn, 'a set, 'b set] => 'b set ("(3\\_\\_./ _)" 10) - Union :: (('a set) set) => 'a set ("\\_" [90] 90) - Inter :: (('a set) set) => 'a set ("\\_" [90] 90) - "_Ball" :: [pttrn, 'a set, bool] => bool ("(3\\_\\_./ _)" [0, 0, 10] 10) - "_Bex" :: [pttrn, 'a set, bool] => bool ("(3\\_\\_./ _)" [0, 0, 10] 10) + "_setle" :: "'a set => 'a set => bool" ("op \") + "_setle" :: "'a set => 'a set => bool" ("(_/ \ _)" [50, 51] 50) + "_setless" :: "'a set => 'a set => bool" ("op \") + "_setless" :: "'a set => 'a set => bool" ("(_/ \ _)" [50, 51] 50) + "op Int" :: "'a set => 'a set => 'a set" (infixl "\" 70) + "op Un" :: "'a set => 'a set => 'a set" (infixl "\" 65) + "op :" :: "'a => 'a set => bool" ("op \") + "op :" :: "'a => 'a set => bool" ("(_/ \ _)" [50, 51] 50) + "op ~:" :: "'a => 'a set => bool" ("op \") + "op ~:" :: "'a => 'a set => bool" ("(_/ \ _)" [50, 51] 50) + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" 10) + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" 10) + Union :: "'a set set => 'a set" ("\_" [90] 90) + Inter :: "'a set set => 'a set" ("\_" [90] 90) + "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) translations - "op \\" => "op <= :: [_ set, _ set] => bool" - "op \\" => "op < :: [_ set, _ set] => bool" - + "op \" => "op <= :: _ set => _ set => bool" + "op \" => "op < :: _ set => _ set => bool" -(** Rules and definitions **) +typed_print_translation {* + let + fun le_tr' _ (Type ("fun", (Type ("set", _) :: _))) ts = + list_comb (Syntax.const "_setle", ts) + | le_tr' _ _ _ = raise Match; + + fun less_tr' _ (Type ("fun", (Type ("set", _) :: _))) ts = + list_comb (Syntax.const "_setless", ts) + | less_tr' _ _ _ = raise Match; + in [("op <=", le_tr'), ("op <", less_tr')] end +*} -local +text {* + \medskip Translate between @{text "{e | x1...xn. P}"} and @{text + "{u. EX x1..xn. u = e & P}"}; @{text "{y. EX x1..xn. y = e & P}"} is + only translated if @{text "[0..n] subset bvs(e)"}. +*} + +parse_translation {* + let + val ex_tr = snd (mk_binder_tr ("EX ", "Ex")); -rules + fun nvars (Const ("_idts", _) $ _ $ idts) = nvars idts + 1 + | nvars _ = 1; + + fun setcompr_tr [e, idts, b] = + let + val eq = Syntax.const "op =" $ Bound (nvars idts) $ e; + val P = Syntax.const "op &" $ eq $ b; + val exP = ex_tr [idts, P]; + in Syntax.const "Collect" $ Abs ("", dummyT, exP) end; + + in [("@SetCompr", setcompr_tr)] end; +*} - (* Isomorphisms between Predicates and Sets *) +print_translation {* + let + val ex_tr' = snd (mk_binder_tr' ("Ex", "DUMMY")); + + fun setcompr_tr' [Abs (_, _, P)] = + let + fun check (Const ("Ex", _) $ Abs (_, _, P), n) = check (P, n + 1) + | check (Const ("op &", _) $ (Const ("op =", _) $ Bound m $ e) $ P, n) = + if n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso + ((0 upto (n - 1)) subset add_loose_bnos (e, 0, [])) then () + else raise Match; - mem_Collect_eq "(a : {x. P(x)}) = P(a)" - Collect_mem_eq "{x. x:A} = A" + fun tr' (_ $ abs) = + let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' [abs] + in Syntax.const "@SetCompr" $ e $ idts $ Q end; + in check (P, 0); tr' P end; + in [("Collect", setcompr_tr')] end; +*} + + +subsection {* Rules and definitions *} + +text {* Isomorphisms between predicates and sets. *} +axioms + mem_Collect_eq [iff]: "(a : {x. P(x)}) = P(a)" + Collect_mem_eq [simp]: "{x. x:A} = A" + +defs + Ball_def: "Ball A P == ALL x. x:A --> P(x)" + Bex_def: "Bex A P == EX x. x:A & P(x)" + +defs (overloaded) + subset_def: "A <= B == ALL x:A. x:B" + psubset_def: "A < B == (A::'a set) <= B & ~ A=B" + Compl_def: "- A == {x. ~x:A}" defs - Ball_def "Ball A P == ! x. x:A --> P(x)" - Bex_def "Bex A P == ? x. x:A & P(x)" - subset_def "A <= B == ! x:A. x:B" - psubset_def "A < B == (A::'a set) <= B & ~ A=B" - Compl_def "- A == {x. ~x:A}" - Un_def "A Un B == {x. x:A | x:B}" - Int_def "A Int B == {x. x:A & x:B}" - set_diff_def "A - B == {x. x:A & ~x:B}" - INTER_def "INTER A B == {y. ! x:A. y: B(x)}" - UNION_def "UNION A B == {y. ? x:A. y: B(x)}" - Inter_def "Inter S == (INT x:S. x)" - Union_def "Union S == (UN x:S. x)" - Pow_def "Pow A == {B. B <= A}" - empty_def "{} == {x. False}" - UNIV_def "UNIV == {x. True}" - insert_def "insert a B == {x. x=a} Un B" - image_def "f`A == {y. ? x:A. y=f(x)}" + Un_def: "A Un B == {x. x:A | x:B}" + Int_def: "A Int B == {x. x:A & x:B}" + set_diff_def: "A - B == {x. x:A & ~x:B}" + INTER_def: "INTER A B == {y. ALL x:A. y: B(x)}" + UNION_def: "UNION A B == {y. EX x:A. y: B(x)}" + Inter_def: "Inter S == (INT x:S. x)" + Union_def: "Union S == (UN x:S. x)" + Pow_def: "Pow A == {B. B <= A}" + empty_def: "{} == {x. False}" + UNIV_def: "UNIV == {x. True}" + insert_def: "insert a B == {x. x=a} Un B" + image_def: "f`A == {y. EX x:A. y = f(x)}" + + +subsection {* Lemmas and proof tool setup *} + +subsubsection {* Relating predicates and sets *} + +lemma CollectI [intro!]: "P(a) ==> a : {x. P(x)}" + by simp + +lemma CollectD: "a : {x. P(x)} ==> P(a)" + by simp + +lemma set_ext: "(!!x. (x:A) = (x:B)) ==> A = B" +proof - + case rule_context + show ?thesis + apply (rule prems [THEN ext, THEN arg_cong, THEN box_equals]) + apply (rule Collect_mem_eq) + apply (rule Collect_mem_eq) + done +qed + +lemma Collect_cong: "(!!x. P x = Q x) ==> {x. P(x)} = {x. Q(x)}" + by simp + +lemmas CollectE [elim!] = CollectD [elim_format] + + +subsubsection {* Bounded quantifiers *} + +lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x" + by (simp add: Ball_def) + +lemmas strip = impI allI ballI + +lemma bspec [dest?]: "ALL x:A. P x ==> x:A ==> P x" + by (simp add: Ball_def) + +lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q" + by (unfold Ball_def) blast + +text {* + \medskip This tactic takes assumptions @{prop "ALL x:A. P x"} and + @{prop "a:A"}; creates assumption @{prop "P a"}. +*} + +ML {* + local val ballE = thm "ballE" + in fun ball_tac i = etac ballE i THEN contr_tac (i + 1) end; +*} + +text {* + Gives better instantiation for bound: +*} + +ML_setup {* + claset_ref() := claset() addbefore ("bspec", datac (thm "bspec") 1); +*} + +lemma bexI [intro]: "P x ==> x:A ==> EX x:A. P x" + -- {* Normally the best argument order: @{prop "P x"} constrains the + choice of @{prop "x:A"}. *} + by (unfold Bex_def) blast + +lemma rev_bexI: "x:A ==> P x ==> EX x:A. P x" + -- {* The best argument order when there is only one @{prop "x:A"}. *} + by (unfold Bex_def) blast + +lemma bexCI: "(ALL x:A. ~P x ==> P a) ==> a:A ==> EX x:A. P x" + by (unfold Bex_def) blast + +lemma bexE [elim!]: "EX x:A. P x ==> (!!x. x:A ==> P x ==> Q) ==> Q" + by (unfold Bex_def) blast + +lemma ball_triv [simp]: "(ALL x:A. P) = ((EX x. x:A) --> P)" + -- {* Trival rewrite rule. *} + by (simp add: Ball_def) + +lemma bex_triv [simp]: "(EX x:A. P) = ((EX x. x:A) & P)" + -- {* Dual form for existentials. *} + by (simp add: Bex_def) + +lemma bex_triv_one_point1 [simp]: "(EX x:A. x = a) = (a:A)" + by blast + +lemma bex_triv_one_point2 [simp]: "(EX x:A. a = x) = (a:A)" + by blast + +lemma bex_one_point1 [simp]: "(EX x:A. x = a & P x) = (a:A & P a)" + by blast + +lemma bex_one_point2 [simp]: "(EX x:A. a = x & P x) = (a:A & P a)" + by blast + +lemma ball_one_point1 [simp]: "(ALL x:A. x = a --> P x) = (a:A --> P a)" + by blast + +lemma ball_one_point2 [simp]: "(ALL x:A. a = x --> P x) = (a:A --> P a)" + by blast + +ML_setup {* + let + val Ball_def = thm "Ball_def"; + val Bex_def = thm "Bex_def"; + + val ex_pattern = Thm.read_cterm (Theory.sign_of (the_context ())) + ("EX x:A. P x & Q x", HOLogic.boolT); + + val prove_bex_tac = + rewrite_goals_tac [Bex_def] THEN Quantifier1.prove_one_point_ex_tac; + val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac; + + val all_pattern = Thm.read_cterm (Theory.sign_of (the_context ())) + ("ALL x:A. P x --> Q x", HOLogic.boolT); + + val prove_ball_tac = + rewrite_goals_tac [Ball_def] THEN Quantifier1.prove_one_point_all_tac; + val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac; + + val defBEX_regroup = mk_simproc "defined BEX" [ex_pattern] rearrange_bex; + val defBALL_regroup = mk_simproc "defined BALL" [all_pattern] rearrange_ball; + in + Addsimprocs [defBALL_regroup, defBEX_regroup] + end; +*} + + +subsubsection {* Congruence rules *} + +lemma ball_cong [cong]: + "A = B ==> (!!x. x:B ==> P x = Q x) ==> + (ALL x:A. P x) = (ALL x:B. Q x)" + by (simp add: Ball_def) + +lemma bex_cong [cong]: + "A = B ==> (!!x. x:B ==> P x = Q x) ==> + (EX x:A. P x) = (EX x:B. Q x)" + by (simp add: Bex_def cong: conj_cong) -end +subsubsection {* Subsets *} + +lemma subsetI [intro!]: "(!!x. x:A ==> x:B) ==> A <= B" + by (simp add: subset_def) + +text {* + \medskip Map the type @{text "'a set => anything"} to just @{typ + 'a}; for overloading constants whose first argument has type @{typ + "'a set"}. +*} + +ML {* + fun overload_1st_set s = Blast.overloaded (s, HOLogic.dest_setT o domain_type); +*} + +ML " + (* While (:) is not, its type must be kept + for overloading of = to work. *) + Blast.overloaded (\"op :\", domain_type); + + overload_1st_set \"Ball\"; (*need UNION, INTER also?*) + overload_1st_set \"Bex\"; + + (*Image: retain the type of the set being expressed*) + Blast.overloaded (\"image\", domain_type); +" + +lemma subsetD [elim]: "A <= B ==> c:A ==> c:B" + -- {* Rule in Modus Ponens style. *} + by (unfold subset_def) blast + +declare subsetD [intro?] -- FIXME + +lemma rev_subsetD: "c:A ==> A <= B ==> c:B" + -- {* The same, with reversed premises for use with @{text erule} -- + cf @{text rev_mp}. *} + by (rule subsetD) + +declare rev_subsetD [intro?] -- FIXME + +text {* + \medskip Converts @{prop "A <= B"} to @{prop "x:A ==> x:B"}. +*} + +ML {* + local val rev_subsetD = thm "rev_subsetD" + in fun impOfSubs th = th RSN (2, rev_subsetD) end; +*} + +lemma subsetCE [elim]: "A <= B ==> (c~:A ==> P) ==> (c:B ==> P) ==> P" + -- {* Classical elimination rule. *} + by (unfold subset_def) blast + +text {* + \medskip Takes assumptions @{prop "A <= B"}; @{prop "c:A"} and + creates the assumption @{prop "c:B"}. +*} + +ML {* + local val subsetCE = thm "subsetCE" + in fun set_mp_tac i = etac subsetCE i THEN mp_tac i end; +*} + +lemma contra_subsetD: "A <= B ==> c ~: B ==> c ~: A" + by blast + +lemma subset_refl: "A <= (A::'a set)" + by fast + +lemma subset_trans: "A <= B ==> B <= C ==> A <= (C::'a set)" + by blast -ML +subsubsection {* Equality *} + +lemma subset_antisym [intro!]: "A <= B ==> B <= A ==> A = (B::'a set)" + -- {* Anti-symmetry of the subset relation. *} + by (rule set_ext) (blast intro: subsetD) + +lemmas equalityI = subset_antisym + +text {* + \medskip Equality rules from ZF set theory -- are they appropriate + here? +*} + +lemma equalityD1: "A = B ==> A <= (B::'a set)" + by (simp add: subset_refl) + +lemma equalityD2: "A = B ==> B <= (A::'a set)" + by (simp add: subset_refl) + +text {* + \medskip Be careful when adding this to the claset as @{text + subset_empty} is in the simpset: @{prop "A = {}"} goes to @{prop "{} + <= A"} and @{prop "A <= {}"} and then back to @{prop "A = {}"}! +*} + +lemma equalityE: "A = B ==> (A <= B ==> B <= (A::'a set) ==> P) ==> P" + by (simp add: subset_refl) -local +lemma equalityCE [elim]: + "A = B ==> (c:A ==> c:B ==> P) ==> (c~:A ==> c~:B ==> P) ==> P" + by blast + +text {* + \medskip Lemma for creating induction formulae -- for "pattern + matching" on @{text p}. To make the induction hypotheses usable, + apply @{text spec} or @{text bspec} to put universal quantifiers over the free + variables in @{text p}. +*} + +lemma setup_induction: "p:A ==> (!!z. z:A ==> p = z --> R) ==> R" + by simp -(* Set inclusion *) +lemma eqset_imp_iff: "A = B ==> (x : A) = (x : B)" + by simp + + +subsubsection {* The universal set -- UNIV *} + +lemma UNIV_I [simp]: "x : UNIV" + by (simp add: UNIV_def) + +declare UNIV_I [intro] -- {* unsafe makes it less likely to cause problems *} + +lemma UNIV_witness [intro?]: "EX x. x : UNIV" + by simp + +lemma subset_UNIV: "A <= UNIV" + by (rule subsetI) (rule UNIV_I) -fun le_tr' _ (*op <=*) (Type ("fun", (Type ("set", _) :: _))) ts = - list_comb (Syntax.const "_setle", ts) - | le_tr' _ (*op <=*) _ _ = raise Match; +text {* + \medskip Eta-contracting these two rules (to remove @{text P}) + causes them to be ignored because of their interaction with + congruence rules. +*} + +lemma ball_UNIV [simp]: "Ball UNIV P = All P" + by (simp add: Ball_def) + +lemma bex_UNIV [simp]: "Bex UNIV P = Ex P" + by (simp add: Bex_def) + + +subsubsection {* The empty set *} + +lemma empty_iff [simp]: "(c : {}) = False" + by (simp add: empty_def) + +lemma emptyE [elim!]: "a : {} ==> P" + by simp + +lemma empty_subsetI [iff]: "{} <= A" + -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *} + by blast + +lemma equals0I: "(!!y. y:A ==> False) ==> A = {}" + by blast -fun less_tr' _ (*op <*) (Type ("fun", (Type ("set", _) :: _))) ts = - list_comb (Syntax.const "_setless", ts) - | less_tr' _ (*op <*) _ _ = raise Match; +lemma equals0D: "A={} ==> a ~: A" + -- {* Use for reasoning about disjointness: @{prop "A Int B = {}"} *} + by blast + +lemma ball_empty [simp]: "Ball {} P = True" + by (simp add: Ball_def) + +lemma bex_empty [simp]: "Bex {} P = False" + by (simp add: Bex_def) + +lemma UNIV_not_empty [iff]: "UNIV ~= {}" + by (blast elim: equalityE) + + +section {* The Powerset operator -- Pow *} + +lemma Pow_iff [iff]: "(A : Pow B) = (A <= B)" + by (simp add: Pow_def) + +lemma PowI: "A <= B ==> A : Pow B" + by (simp add: Pow_def) + +lemma PowD: "A : Pow B ==> A <= B" + by (simp add: Pow_def) + +lemma Pow_bottom: "{}: Pow B" + by simp + +lemma Pow_top: "A : Pow A" + by (simp add: subset_refl) -(* Translates between { e | x1..xn. P} and {u. ? x1..xn. u=e & P} *) -(* {y. ? x1..xn. y = e & P} is only translated if [0..n] subset bvs(e) *) +subsubsection {* Set complement *} + +lemma Compl_iff [simp]: "(c : -A) = (c~:A)" + by (unfold Compl_def) blast + +lemma ComplI [intro!]: "(c:A ==> False) ==> c : -A" + by (unfold Compl_def) blast + +text {* + \medskip This form, with negated conclusion, works well with the + Classical prover. Negated assumptions behave like formulae on the + right side of the notional turnstile ... *} + +lemma ComplD: "c : -A ==> c~:A" + by (unfold Compl_def) blast + +lemmas ComplE [elim!] = ComplD [elim_format] + + +subsubsection {* Binary union -- Un *} -val ex_tr = snd(mk_binder_tr("EX ","Ex")); +lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)" + by (unfold Un_def) blast + +lemma UnI1 [elim?]: "c:A ==> c : A Un B" + by simp + +lemma UnI2 [elim?]: "c:B ==> c : A Un B" + by simp -fun nvars(Const("_idts",_) $ _ $ idts) = nvars(idts)+1 - | nvars(_) = 1; +text {* + \medskip Classical introduction rule: no commitment to @{prop A} vs + @{prop B}. +*} + +lemma UnCI [intro!]: "(c~:B ==> c:A) ==> c : A Un B" + by auto + +lemma UnE [elim!]: "c : A Un B ==> (c:A ==> P) ==> (c:B ==> P) ==> P" + by (unfold Un_def) blast + + +section {* Binary intersection -- Int *} -fun setcompr_tr[e,idts,b] = - let val eq = Syntax.const("op =") $ Bound(nvars(idts)) $ e - val P = Syntax.const("op &") $ eq $ b - val exP = ex_tr [idts,P] - in Syntax.const("Collect") $ Abs("",dummyT,exP) end; +lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" + by (unfold Int_def) blast + +lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B" + by simp + +lemma IntD1: "c : A Int B ==> c:A" + by simp + +lemma IntD2: "c : A Int B ==> c:B" + by simp + +lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P" + by simp + + +section {* Set difference *} + +lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)" + by (unfold set_diff_def) blast -val ex_tr' = snd(mk_binder_tr' ("Ex","DUMMY")); +lemma DiffI [intro!]: "c : A ==> c ~: B ==> c : A - B" + by simp + +lemma DiffD1: "c : A - B ==> c : A" + by simp + +lemma DiffD2: "c : A - B ==> c : B ==> P" + by simp + +lemma DiffE [elim!]: "c : A - B ==> (c:A ==> c~:B ==> P) ==> P" + by simp + + +subsubsection {* Augmenting a set -- insert *} + +lemma insert_iff [simp]: "(a : insert b A) = (a = b | a:A)" + by (unfold insert_def) blast + +lemma insertI1: "a : insert a B" + by simp + +lemma insertI2: "a : B ==> a : insert b B" + by simp -fun setcompr_tr'[Abs(_,_,P)] = - let fun ok(Const("Ex",_)$Abs(_,_,P),n) = ok(P,n+1) - | ok(Const("op &",_) $ (Const("op =",_) $ Bound(m) $ e) $ P, n) = - if n>0 andalso m=n andalso not(loose_bvar1(P,n)) andalso - ((0 upto (n-1)) subset add_loose_bnos(e,0,[])) - then () else raise Match +lemma insertE [elim!]: "a : insert b A ==> (a = b ==> P) ==> (a:A ==> P) ==> P" + by (unfold insert_def) blast + +lemma insertCI [intro!]: "(a~:B ==> a = b) ==> a: insert b B" + -- {* Classical introduction rule. *} + by auto + +lemma subset_insert_iff: "(A <= insert x B) = (if x:A then A - {x} <= B else A <= B)" + by auto + + +subsubsection {* Singletons, using insert *} + +lemma singletonI [intro!]: "a : {a}" + -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *} + by (rule insertI1) + +lemma singletonD: "b : {a} ==> b = a" + by blast + +lemmas singletonE [elim!] = singletonD [elim_format] + +lemma singleton_iff: "(b : {a}) = (b = a)" + by blast + +lemma singleton_inject [dest!]: "{a} = {b} ==> a = b" + by blast + +lemma singleton_insert_inj_eq [iff]: "({b} = insert a A) = (a = b & A <= {b})" + by blast + +lemma singleton_insert_inj_eq' [iff]: "(insert a A = {b}) = (a = b & A <= {b})" + by blast + +lemma subset_singletonD: "A <= {x} ==> A={} | A = {x}" + by fast + +lemma singleton_conv [simp]: "{x. x = a} = {a}" + by blast + +lemma singleton_conv2 [simp]: "{x. a = x} = {a}" + by blast - fun tr'(_ $ abs) = - let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr'[abs] - in Syntax.const("@SetCompr") $ e $ idts $ Q end - in ok(P,0); tr'(P) end; +lemma diff_single_insert: "A - {x} <= B ==> x : A ==> A <= insert x B" + by blast + + +subsubsection {* Unions of families *} + +text {* + @{term [source] "UN x:A. B x"} is @{term "Union (B`A)"}. +*} + +lemma UN_iff [simp]: "(b: (UN x:A. B x)) = (EX x:A. b: B x)" + by (unfold UNION_def) blast + +lemma UN_I [intro]: "a:A ==> b: B a ==> b: (UN x:A. B x)" + -- {* The order of the premises presupposes that @{term A} is rigid; + @{term b} may be flexible. *} + by auto + +lemma UN_E [elim!]: "b : (UN x:A. B x) ==> (!!x. x:A ==> b: B x ==> R) ==> R" + by (unfold UNION_def) blast -in +lemma UN_cong [cong]: + "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" + by (simp add: UNION_def) + + +subsubsection {* Intersections of families *} + +text {* @{term [source] "INT x:A. B x"} is @{term "Inter (B`A)"}. *} + +lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" + by (unfold INTER_def) blast -val parse_translation = [("@SetCompr", setcompr_tr)]; -val print_translation = [("Collect", setcompr_tr')]; -val typed_print_translation = [("op <=", le_tr'), ("op <", less_tr')]; +lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)" + by (unfold INTER_def) blast + +lemma INT_D [elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a" + by auto + +lemma INT_E [elim]: "b : (INT x:A. B x) ==> (b: B a ==> R) ==> (a~:A ==> R) ==> R" + -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *} + by (unfold INTER_def) blast + +lemma INT_cong [cong]: + "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)" + by (simp add: INTER_def) -end; +subsubsection {* Union *} + +lemma Union_iff [simp]: "(A : Union C) = (EX X:C. A:X)" + by (unfold Union_def) blast + +lemma UnionI [intro]: "X:C ==> A:X ==> A : Union C" + -- {* The order of the premises presupposes that @{term C} is rigid; + @{term A} may be flexible. *} + by auto + +lemma UnionE [elim!]: "A : Union C ==> (!!X. A:X ==> X:C ==> R) ==> R" + by (unfold Union_def) blast + + +subsubsection {* Inter *} + +lemma Inter_iff [simp]: "(A : Inter C) = (ALL X:C. A:X)" + by (unfold Inter_def) blast + +lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C" + by (simp add: Inter_def) + +text {* + \medskip A ``destruct'' rule -- every @{term X} in @{term C} + contains @{term A} as an element, but @{prop "A:X"} can hold when + @{prop "X:C"} does not! This rule is analogous to @{text spec}. +*} + +lemma InterD [elim]: "A : Inter C ==> X:C ==> A:X" + by auto + +lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R" + -- {* ``Classical'' elimination rule -- does not require proving + @{prop "X:C"}. *} + by (unfold Inter_def) blast + +text {* + \medskip Image of a set under a function. Frequently @{term b} does + not have the syntactic form of @{term "f x"}. +*} + +lemma image_eqI [simp, intro]: "b = f x ==> x:A ==> b : f`A" + by (unfold image_def) blast + +lemma imageI: "x : A ==> f x : f ` A" + by (rule image_eqI) (rule refl) + +lemma rev_image_eqI: "x:A ==> b = f x ==> b : f`A" + -- {* This version's more effective when we already have the + required @{term x}. *} + by (unfold image_def) blast + +lemma imageE [elim!]: + "b : (%x. f x)`A ==> (!!x. b = f x ==> x:A ==> P) ==> P" + -- {* The eta-expansion gives variable-name preservation. *} + by (unfold image_def) blast + +lemma image_Un: "f`(A Un B) = f`A Un f`B" + by blast + +lemma image_iff: "(z : f`A) = (EX x:A. z = f x)" + by blast + +lemma image_subset_iff: "(f`A <= B) = (ALL x:A. f x: B)" + -- {* This rewrite rule would confuse users if made default. *} + by blast + +lemma subset_image_iff: "(B <= f ` A) = (EX AA. AA <= A & B = f ` AA)" + apply safe + prefer 2 apply fast + apply (rule_tac x = "{a. a : A & f a : B}" in exI) + apply fast + done + +lemma image_subsetI: "(!!x. x:A ==> f x : B) ==> f`A <= B" + -- {* Replaces the three steps @{text subsetI}, @{text imageE}, + @{text hypsubst}, but breaks too many existing proofs. *} + by blast + +text {* + \medskip Range of a function -- just a translation for image! +*} + +lemma range_eqI: "b = f x ==> b : range f" + by simp + +lemma rangeI: "f x : range f" + by simp + +lemma rangeE [elim?]: "b : range (%x. f x) ==> (!!x. b = f x ==> P) ==> P" + by blast + + +subsubsection {* Set reasoning tools *} + +text {* + Rewrite rules for boolean case-splitting: faster than @{text + "split_if [split]"}. +*} + +lemma split_if_eq1: "((if Q then x else y) = b) = ((Q --> x = b) & (~ Q --> y = b))" + by (rule split_if) + +lemma split_if_eq2: "(a = (if Q then x else y)) = ((Q --> a = x) & (~ Q --> a = y))" + by (rule split_if) + +text {* + Split ifs on either side of the membership relation. Not for @{text + "[simp]"} -- can cause goals to blow up! +*} + +lemma split_if_mem1: "((if Q then x else y) : b) = ((Q --> x : b) & (~ Q --> y : b))" + by (rule split_if) + +lemma split_if_mem2: "(a : (if Q then x else y)) = ((Q --> a : x) & (~ Q --> a : y))" + by (rule split_if) + +lemmas split_ifs = if_bool_eq_conj split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2 + +lemmas mem_simps = + insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff + mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff + -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} + +(*Would like to add these, but the existing code only searches for the + outer-level constant, which in this case is just "op :"; we instead need + to use term-nets to associate patterns with rules. Also, if a rule fails to + apply, then the formula should be kept. + [("uminus", Compl_iff RS iffD1), ("op -", [Diff_iff RS iffD1]), + ("op Int", [IntD1,IntD2]), + ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] + *) + +ML_setup {* + val mksimps_pairs = [("Ball", [thm "bspec"])] @ mksimps_pairs; + simpset_ref() := simpset() setmksimps (mksimps mksimps_pairs); +*} + +declare subset_UNIV [simp] subset_refl [simp] + + +subsubsection {* The ``proper subset'' relation *} + +lemma psubsetI [intro!]: "(A::'a set) <= B ==> A ~= B ==> A < B" + by (unfold psubset_def) blast + +lemma psubset_insert_iff: + "(A < insert x B) = (if x:B then A < B else if x:A then A - {x} < B else A <= B)" + apply (simp add: psubset_def subset_insert_iff) + apply blast + done + +lemma psubset_eq: "((A::'a set) < B) = (A <= B & A ~= B)" + by (simp only: psubset_def) + +lemma psubset_imp_subset: "(A::'a set) < B ==> A <= B" + by (simp add: psubset_eq) + +lemma psubset_subset_trans: "(A::'a set) < B ==> B <= C ==> A < C" + by (auto simp add: psubset_eq) + +lemma subset_psubset_trans: "(A::'a set) <= B ==> B < C ==> A < C" + by (auto simp add: psubset_eq) + +lemma psubset_imp_ex_mem: "A < B ==> EX b. b : (B - A)" + by (unfold psubset_def) blast + +lemma atomize_ball: + "(!!x. x:A ==> P x) == Trueprop (ALL x:A. P x)" + by (simp only: Ball_def atomize_all atomize_imp) + +declare atomize_ball [symmetric, rulify] + + +subsection {* Further set-theory lemmas *} + +use "subset.ML" +use "equalities.ML" +use "mono.ML" + +end diff -r a394d3e9c8bb -r 0a3dace545c5 src/HOL/Typedef.thy --- a/src/HOL/Typedef.thy Sun Oct 28 22:58:39 2001 +0100 +++ b/src/HOL/Typedef.thy Sun Oct 28 22:59:12 2001 +0100 @@ -3,18 +3,15 @@ Author: Markus Wenzel, TU Munich *) -header {* Set-theory lemmas and HOL type definitions *} +header {* HOL type definitions *} theory Typedef = Set -files "subset.ML" "equalities.ML" "mono.ML" ("Tools/typedef_package.ML"): +files ("Tools/typedef_package.ML"): -(*belongs to theory Set*) -declare atomize_ball [symmetric, rulify] - -(* Courtesy of Stephan Merz *) lemma Least_mono: "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)" + -- {* Courtesy of Stephan Merz *} apply clarify apply (erule_tac P = "%x. x : S" in LeastI2) apply fast diff -r a394d3e9c8bb -r 0a3dace545c5 src/HOL/subset.ML --- a/src/HOL/subset.ML Sun Oct 28 22:58:39 2001 +0100 +++ b/src/HOL/subset.ML Sun Oct 28 22:59:12 2001 +0100 @@ -7,6 +7,163 @@ operations. *) +(* legacy ML bindings *) + +val Ball_def = thm "Ball_def"; +val Bex_def = thm "Bex_def"; +val Collect_mem_eq = thm "Collect_mem_eq"; +val Compl_def = thm "Compl_def"; +val INTER_def = thm "INTER_def"; +val Int_def = thm "Int_def"; +val Inter_def = thm "Inter_def"; +val Pow_def = thm "Pow_def"; +val UNION_def = thm "UNION_def"; +val UNIV_def = thm "UNIV_def"; +val Un_def = thm "Un_def"; +val Union_def = thm "Union_def"; +val empty_def = thm "empty_def"; +val image_def = thm "image_def"; +val insert_def = thm "insert_def"; +val mem_Collect_eq = thm "mem_Collect_eq"; +val psubset_def = thm "psubset_def"; +val set_diff_def = thm "set_diff_def"; +val subset_def = thm "subset_def"; +val CollectI = thm "CollectI"; +val CollectD = thm "CollectD"; +val set_ext = thm "set_ext"; +val Collect_cong = thm "Collect_cong"; +val CollectE = thm "CollectE"; +val ballI = thm "ballI"; +val strip = thms "strip"; +val bspec = thm "bspec"; +val ballE = thm "ballE"; +val bexI = thm "bexI"; +val rev_bexI = thm "rev_bexI"; +val bexCI = thm "bexCI"; +val bexE = thm "bexE"; +val ball_triv = thm "ball_triv"; +val bex_triv = thm "bex_triv"; +val bex_triv_one_point1 = thm "bex_triv_one_point1"; +val bex_triv_one_point2 = thm "bex_triv_one_point2"; +val bex_one_point1 = thm "bex_one_point1"; +val bex_one_point2 = thm "bex_one_point2"; +val ball_one_point1 = thm "ball_one_point1"; +val ball_one_point2 = thm "ball_one_point2"; +val ball_cong = thm "ball_cong"; +val bex_cong = thm "bex_cong"; +val subsetI = thm "subsetI"; +val subsetD = thm "subsetD"; +val rev_subsetD = thm "rev_subsetD"; +val subsetCE = thm "subsetCE"; +val contra_subsetD = thm "contra_subsetD"; +val subset_refl = thm "subset_refl"; +val subset_trans = thm "subset_trans"; +val subset_antisym = thm "subset_antisym"; +val equalityI = thm "equalityI"; +val equalityD1 = thm "equalityD1"; +val equalityD2 = thm "equalityD2"; +val equalityE = thm "equalityE"; +val equalityCE = thm "equalityCE"; +val setup_induction = thm "setup_induction"; +val eqset_imp_iff = thm "eqset_imp_iff"; +val UNIV_I = thm "UNIV_I"; +val UNIV_witness = thm "UNIV_witness"; +val subset_UNIV = thm "subset_UNIV"; +val ball_UNIV = thm "ball_UNIV"; +val bex_UNIV = thm "bex_UNIV"; +val empty_iff = thm "empty_iff"; +val emptyE = thm "emptyE"; +val empty_subsetI = thm "empty_subsetI"; +val equals0I = thm "equals0I"; +val equals0D = thm "equals0D"; +val ball_empty = thm "ball_empty"; +val bex_empty = thm "bex_empty"; +val UNIV_not_empty = thm "UNIV_not_empty"; +val Pow_iff = thm "Pow_iff"; +val PowI = thm "PowI"; +val PowD = thm "PowD"; +val Pow_bottom = thm "Pow_bottom"; +val Pow_top = thm "Pow_top"; +val Compl_iff = thm "Compl_iff"; +val ComplI = thm "ComplI"; +val ComplD = thm "ComplD"; +val ComplE = thm "ComplE"; +val Un_iff = thm "Un_iff"; +val UnI1 = thm "UnI1"; +val UnI2 = thm "UnI2"; +val UnCI = thm "UnCI"; +val UnE = thm "UnE"; +val Int_iff = thm "Int_iff"; +val IntI = thm "IntI"; +val IntD1 = thm "IntD1"; +val IntD2 = thm "IntD2"; +val IntE = thm "IntE"; +val Diff_iff = thm "Diff_iff"; +val DiffI = thm "DiffI"; +val DiffD1 = thm "DiffD1"; +val DiffD2 = thm "DiffD2"; +val DiffE = thm "DiffE"; +val insert_iff = thm "insert_iff"; +val insertI1 = thm "insertI1"; +val insertI2 = thm "insertI2"; +val insertE = thm "insertE"; +val insertCI = thm "insertCI"; +val subset_insert_iff = thm "subset_insert_iff"; +val singletonI = thm "singletonI"; +val singletonD = thm "singletonD"; +val singletonE = thm "singletonE"; +val singleton_iff = thm "singleton_iff"; +val singleton_inject = thm "singleton_inject"; +val singleton_insert_inj_eq = thm "singleton_insert_inj_eq"; +val singleton_insert_inj_eq' = thm "singleton_insert_inj_eq'"; +val subset_singletonD = thm "subset_singletonD"; +val singleton_conv = thm "singleton_conv"; +val singleton_conv2 = thm "singleton_conv2"; +val diff_single_insert = thm "diff_single_insert"; +val UN_iff = thm "UN_iff"; +val UN_I = thm "UN_I"; +val UN_E = thm "UN_E"; +val UN_cong = thm "UN_cong"; +val INT_iff = thm "INT_iff"; +val INT_I = thm "INT_I"; +val INT_D = thm "INT_D"; +val INT_E = thm "INT_E"; +val INT_cong = thm "INT_cong"; +val Union_iff = thm "Union_iff"; +val UnionI = thm "UnionI"; +val UnionE = thm "UnionE"; +val Inter_iff = thm "Inter_iff"; +val InterI = thm "InterI"; +val InterD = thm "InterD"; +val InterE = thm "InterE"; +val image_eqI = thm "image_eqI"; +val imageI = thm "imageI"; +val rev_image_eqI = thm "rev_image_eqI"; +val imageE = thm "imageE"; +val image_Un = thm "image_Un"; +val image_iff = thm "image_iff"; +val image_subset_iff = thm "image_subset_iff"; +val subset_image_iff = thm "subset_image_iff"; +val image_subsetI = thm "image_subsetI"; +val range_eqI = thm "range_eqI"; +val rangeI = thm "rangeI"; +val rangeE = thm "rangeE"; +val split_if_eq1 = thm "split_if_eq1"; +val split_if_eq2 = thm "split_if_eq2"; +val split_if_mem1 = thm "split_if_mem1"; +val split_if_mem2 = thm "split_if_mem2"; +val split_ifs = thms "split_ifs"; +val mem_simps = thms "mem_simps"; +val psubsetI = thm "psubsetI"; +val psubset_insert_iff = thm "psubset_insert_iff"; +val psubset_eq = thm "psubset_eq"; +val psubset_imp_subset = thm "psubset_imp_subset"; +val psubset_subset_trans = thm "psubset_subset_trans"; +val subset_psubset_trans = thm "subset_psubset_trans"; +val psubset_imp_ex_mem = thm "psubset_imp_ex_mem"; +val atomize_ball = thm "atomize_ball"; + + (*** insert ***) Goal "B <= insert a B";