# HG changeset patch # User lcp # Date 776966650 -7200 # Node ID 98b88551e10251f310841ff5da5f67b4e4a210ef # Parent 4530c45370b431f56f56d1121aa147562a77cbe0 ZF/func/empty_fun: renamed from fun_empty ZF/func/single_fun: replaces the weaker fun_single ZF/func/fun_single_lemma: deleted ZF/func.thy: now depends upon equalities.thy diff -r 4530c45370b4 -r 98b88551e102 src/ZF/func.ML --- a/src/ZF/func.ML Mon Aug 15 16:12:35 1994 +0200 +++ b/src/ZF/func.ML Mon Aug 15 18:04:10 1994 +0200 @@ -49,13 +49,23 @@ (*Empty function spaces*) goalw ZF.thy [Pi_def] "Pi(0,A) = {0}"; -by (fast_tac (ZF_cs addIs [equalityI]) 1); +by (fast_tac eq_cs 1); val Pi_empty1 = result(); goalw ZF.thy [Pi_def] "!!A a. a:A ==> A->0 = 0"; -by (fast_tac (ZF_cs addIs [equalityI]) 1); +by (fast_tac eq_cs 1); val Pi_empty2 = result(); +(*The empty function*) +goal ZF.thy "0: 0->A"; +by (fast_tac (ZF_cs addIs [PiI]) 1); +val empty_fun = result(); + +(*The singleton function*) +goalw ZF.thy [Pi_def] "{} : {a} -> {b}"; +by (fast_tac eq_cs 1); +val single_fun = result(); + (*** Function Application ***) goal ZF.thy "!!a b f. [| : f; : f; f: Pi(A,B) |] ==> b=c"; @@ -357,17 +367,11 @@ (*** Extensions of functions ***) -(*Singleton function -- in the underlying form of singletons*) -goal ZF.thy "Upair(,) : Upair(a,a) -> Upair(b,b)"; -by (fast_tac (ZF_cs addIs [PiI]) 1); -val fun_single_lemma = result(); - -goalw ZF.thy [cons_def] +goal ZF.thy "!!f A B. [| f: A->B; c~:A |] ==> cons(,f) : cons(c,A) -> cons(b,B)"; -by (rtac (fun_single_lemma RS fun_disjoint_Un) 1); -by (assume_tac 1); -by (rtac equals0I 1); -by (fast_tac ZF_cs 1); +by (forward_tac [single_fun RS fun_disjoint_Un] 1); +by (asm_full_simp_tac (FOL_ss addsimps [cons_eq]) 2); +by (fast_tac eq_cs 1); val fun_extend = result(); goal ZF.thy "!!f A B. [| f: A->B; a:A; c~:A |] ==> cons(,f)`a = f`a"; @@ -382,16 +386,6 @@ by (REPEAT (assume_tac 1)); val fun_extend_apply2 = result(); -(*The empty function*) -goal ZF.thy "0: 0->A"; -by (fast_tac (ZF_cs addIs [PiI]) 1); -val fun_empty = result(); - -(*The singleton function*) -goal ZF.thy "{} : {a} -> cons(b,C)"; -by (REPEAT (ares_tac [fun_extend,fun_empty,notI] 1 ORELSE etac emptyE 1)); -val fun_single = result(); - goal ZF.thy "!!c. c ~: A ==> cons(c,A) -> B = (UN f: A->B. UN b:B. {cons(, f)})"; by (safe_tac eq_cs); diff -r 4530c45370b4 -r 98b88551e102 src/ZF/func.thy --- a/src/ZF/func.thy Mon Aug 15 16:12:35 1994 +0200 +++ b/src/ZF/func.thy Mon Aug 15 18:04:10 1994 +0200 @@ -1,3 +1,3 @@ (*Dummy theory to document dependencies *) -func = "domrange" \ No newline at end of file +func = "domrange" + "equalities"