src/ZF/Perm.ML
changeset 760 f0200e91b272
parent 735 f99621230c2e
child 782 200a16083201
--- a/src/ZF/Perm.ML	Wed Dec 07 12:34:47 1994 +0100
+++ b/src/ZF/Perm.ML	Wed Dec 07 13:12:04 1994 +0100
@@ -15,16 +15,16 @@
 
 goalw Perm.thy [surj_def] "!!f A B. f: surj(A,B) ==> f: A->B";
 by (etac CollectD1 1);
-val surj_is_fun = result();
+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] 
 		    addEs [range_of_fun,domain_type]) 1);
-val fun_is_surj = result();
+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);
-val surj_range = result();
+qed "surj_range";
 
 (** A function with a right inverse is a surjection **)
 
@@ -32,7 +32,7 @@
     "[| 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);
-val f_imp_surjective = result();
+qed "f_imp_surjective";
 
 val prems = goal Perm.thy
     "[| !!x. x:A ==> c(x): B;		\
@@ -41,27 +41,27 @@
 \    |] ==> (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)) ));
-val lam_surjective = result();
+qed "lam_surjective";
 
 (*Cantor's theorem revisited*)
 goalw Perm.thy [surj_def] "f ~: surj(A,Pow(A))";
 by (safe_tac ZF_cs);
 by (cut_facts_tac [cantor] 1);
 by (fast_tac subset_cs 1);
-val cantor_surj = result();
+qed "cantor_surj";
 
 
 (** Injective function space **)
 
 goalw Perm.thy [inj_def] "!!f A B. f: inj(A,B) ==> f: A->B";
 by (etac CollectD1 1);
-val inj_is_fun = result();
+qed "inj_is_fun";
 
 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);
-val inj_equality = result();
+qed "inj_equality";
 
 (** A function with a left inverse is an injection **)
 
@@ -71,7 +71,7 @@
 by (safe_tac ZF_cs);
 by (eresolve_tac [subst_context RS box_equals] 1);
 by (REPEAT (ares_tac prems 1));
-val f_imp_injective = result();
+qed "f_imp_injective";
 
 val prems = goal Perm.thy
     "[| !!x. x:A ==> c(x): B;		\
@@ -79,17 +79,17 @@
 \    |] ==> (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)) ));
-val lam_injective = result();
+qed "lam_injective";
 
 (** Bijections **)
 
 goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: inj(A,B)";
 by (etac IntD1 1);
-val bij_is_inj = result();
+qed "bij_is_inj";
 
 goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> f: surj(A,B)";
 by (etac IntD2 1);
-val bij_is_surj = result();
+qed "bij_is_surj";
 
 (* f: bij(A,B) ==> f: A->B *)
 val bij_is_fun = standard (bij_is_inj RS inj_is_fun);
@@ -101,46 +101,46 @@
 \       !!y. y:B ==> c(d(y)) = y	\
 \    |] ==> (lam x:A.c(x)) : bij(A,B)";
 by (REPEAT (ares_tac (prems @ [IntI, lam_injective, lam_surjective]) 1));
-val lam_bijective = result();
+qed "lam_bijective";
 
 
 (** Identity function **)
 
 val [prem] = goalw Perm.thy [id_def] "a:A ==> <a,a> : id(A)";  
 by (rtac (prem RS lamI) 1);
-val idI = result();
+qed "idI";
 
 val major::prems = goalw Perm.thy [id_def]
     "[| p: id(A);  !!x.[| x:A; p=<x,x> |] ==> P  \
 \    |] ==>  P";  
 by (rtac (major RS lamE) 1);
 by (REPEAT (ares_tac prems 1));
-val idE = result();
+qed "idE";
 
 goalw Perm.thy [id_def] "id(A) : A->A";  
 by (rtac lam_type 1);
 by (assume_tac 1);
-val id_type = result();
+qed "id_type";
 
 val [prem] = goalw Perm.thy [id_def] "A<=B ==> id(A) <= id(B)";
 by (rtac (prem RS lam_mono) 1);
-val id_mono = result();
+qed "id_mono";
 
 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);
-val id_subset_inj = result();
+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);
-val id_surj = result();
+qed "id_surj";
 
 goalw Perm.thy [bij_def] "id(A): bij(A,A)";
 by (fast_tac (ZF_cs addIs [id_inj,id_surj]) 1);
-val id_bij = result();
+qed "id_bij";
 
 goalw Perm.thy [id_def] "A <= B <-> id(A) : A->B";
 by (safe_tac ZF_cs);
@@ -148,7 +148,7 @@
 by (dtac apply_type 1);
 by (assume_tac 1);
 by (asm_full_simp_tac ZF_ss 1);
-val subset_iff_id = result();
+qed "subset_iff_id";
 
 
 (*** Converse of a function ***)
@@ -160,7 +160,7 @@
 by (deepen_tac ZF_cs 0 2);
 by (simp_tac (ZF_ss addsimps [function_def, converse_iff]) 1);
 by (fast_tac (ZF_cs addEs [prem RSN (3,inj_equality)]) 1);
-val inj_converse_fun = result();
+qed "inj_converse_fun";
 
 (** Equations for converse(f) **)
 
@@ -168,12 +168,12 @@
 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);
-val left_inverse_lemma = result();
+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);
-val left_inverse = result();
+qed "left_inverse";
 
 val left_inverse_bij = bij_is_inj RS left_inverse;
 
@@ -181,21 +181,21 @@
     "[| f: A->B;  converse(f): C->A;  b: C |] ==> f`(converse(f)`b) = b";
 by (rtac (apply_Pair RS (converseD RS apply_equality)) 1);
 by (REPEAT (resolve_tac prems 1));
-val right_inverse_lemma = result();
+qed "right_inverse_lemma";
 
 (*Should the premises be f:surj(A,B), b:B for symmetry with left_inverse?
   No: they would not imply that converse(f) was a function! *)
 goal Perm.thy "!!f. [| f: inj(A,B);  b: range(f) |] ==> f`(converse(f)`b) = b";
 by (rtac right_inverse_lemma 1);
 by (REPEAT (ares_tac [inj_converse_fun,inj_is_fun] 1));
-val right_inverse = result();
+qed "right_inverse";
 
 goalw Perm.thy [bij_def]
     "!!f. [| f: bij(A,B);  b: B |] ==> f`(converse(f)`b) = b";
 by (EVERY1 [etac IntE, etac right_inverse, 
 	    etac (surj_range RS ssubst),
 	    assume_tac]);
-val right_inverse_bij = result();
+qed "right_inverse_bij";
 
 (** Converses of injections, surjections, bijections **)
 
@@ -204,13 +204,13 @@
 by (eresolve_tac [inj_converse_fun] 1);
 by (resolve_tac [right_inverse] 1);
 by (REPEAT (assume_tac 1));
-val inj_converse_inj = result();
+qed "inj_converse_inj";
 
 goal Perm.thy "!!f A B. f: inj(A,B) ==> converse(f): surj(range(f), A)";
 by (REPEAT (ares_tac [f_imp_surjective, inj_converse_fun] 1));
 by (REPEAT (ares_tac [left_inverse] 2));
 by (REPEAT (ares_tac [inj_is_fun, range_of_fun RS apply_type] 1));
-val inj_converse_surj = result();
+qed "inj_converse_surj";
 
 goalw Perm.thy [bij_def] "!!f A B. f: bij(A,B) ==> converse(f): bij(B,A)";
 by (etac IntE 1);
@@ -218,7 +218,7 @@
 by (rtac IntI 1);
 by (etac inj_converse_inj 1);
 by (etac inj_converse_surj 1);
-val bij_converse_bij = result();
+qed "bij_converse_bij";
 
 
 (** Composition of two relations **)
@@ -227,7 +227,7 @@
 
 goalw Perm.thy [comp_def] "!!r s. [| <a,b>:s; <b,c>:r |] ==> <a,c> : r O s";
 by (fast_tac ZF_cs 1);
-val compI = result();
+qed "compI";
 
 val prems = goalw Perm.thy [comp_def]
     "[| xz : r O s;  \
@@ -235,7 +235,7 @@
 \    |] ==> P";
 by (cut_facts_tac prems 1);
 by (REPEAT (eresolve_tac [CollectE, exE, conjE] 1 ORELSE ares_tac prems 1));
-val compE = result();
+qed "compE";
 
 val compEpair = 
     rule_by_tactic (REPEAT_FIRST (etac Pair_inject ORELSE' bound_hyp_subst_tac)
@@ -249,56 +249,56 @@
 (*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);
-val range_comp = result();
+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);
-val range_comp_eq = result();
+qed "range_comp_eq";
 
 goal Perm.thy "domain(r O s) <= domain(s)";
 by (fast_tac comp_cs 1);
-val domain_comp = result();
+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);
-val domain_comp_eq = result();
+qed "domain_comp_eq";
 
 goal Perm.thy "(r O s)``A = r``(s``A)";
 by (fast_tac (comp_cs addIs [equalityI]) 1);
-val image_comp = result();
+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);
-val comp_mono = result();
+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);
-val comp_rel = result();
+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);
-val comp_assoc = result();
+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);
-val left_comp_id = result();
+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);
-val right_comp_id = result();
+qed "right_comp_id";
 
 
 (** Composition preserves functions, injections, and surjections **)
@@ -306,7 +306,7 @@
 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);
-val comp_function = result();
+qed "comp_function";
 
 goalw Perm.thy [Pi_def]
     "!!f g. [| g: A->B;  f: B->C |] ==> (f O g) : A->C";
@@ -315,12 +315,12 @@
 by (rtac (range_rel_subset RS domain_comp_eq RS ssubst) 2 THEN assume_tac 3);
 by (fast_tac ZF_cs 2);
 by (asm_simp_tac (ZF_ss addsimps [comp_rel]) 1);
-val comp_fun = result();
+qed "comp_fun";
 
 goal Perm.thy "!!f g. [| g: A->B;  f: B->C;  a:A |] ==> (f O g)`a = f`(g`a)";
 by (REPEAT (ares_tac [comp_fun,apply_equality,compI,
 		      apply_Pair,apply_type] 1));
-val comp_fun_apply = result();
+qed "comp_fun_apply";
 
 goal Perm.thy "!!f g. [| g: inj(A,B);  f: inj(B,C) |] ==> (f O g) : inj(A,C)";
 by (res_inst_tac [("d", "%y. converse(g) ` (converse(f) ` y)")]
@@ -328,17 +328,17 @@
 by (REPEAT (ares_tac [comp_fun, inj_is_fun] 1));
 by (asm_simp_tac (ZF_ss addsimps [comp_fun_apply, left_inverse] 
                         setsolver type_auto_tac [inj_is_fun, apply_type]) 1);
-val comp_inj = result();
+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);
-val comp_surj = result();
+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);
-val comp_bij = result();
+qed "comp_bij";
 
 
 (** Dual properties of inj and surj -- useful for proofs from
@@ -350,7 +350,7 @@
 by (safe_tac comp_cs);
 by (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp] 1));
 by (asm_simp_tac (FOL_ss addsimps [comp_fun_apply]) 1);
-val comp_mem_injD1 = result();
+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)";
@@ -362,17 +362,17 @@
 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);
-val comp_mem_injD2 = result();
+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);
-val comp_mem_surjD1 = result();
+qed "comp_mem_surjD1";
 
 goal Perm.thy
     "!!f g. [| (f O g)`a = c;  g: A->B;  f: B->C;  a:A |] ==> f`(g`a) = c";
 by (REPEAT (ares_tac [comp_fun_apply RS sym RS trans] 1));
-val comp_fun_applyD = result();
+qed "comp_fun_applyD";
 
 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)";
@@ -380,7 +380,7 @@
 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);
-val comp_mem_surjD2 = result();
+qed "comp_mem_surjD2";
 
 
 (** inverses of composition **)
@@ -393,7 +393,7 @@
 by (cut_facts_tac [prem RS inj_is_fun] 1);
 by (fast_tac (comp_cs addIs [equalityI,apply_Pair] 
 		      addEs [domain_type, make_elim injfD]) 1);
-val left_comp_inverse = result();
+qed "left_comp_inverse";
 
 (*right inverse of composition; one inclusion is
 	        f: A->B ==> f O converse(f) <= id(B) 
@@ -405,7 +405,7 @@
 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);
-val right_comp_inverse = result();
+qed "right_comp_inverse";
 
 (** Proving that a function is a bijection **)
 
@@ -418,7 +418,7 @@
 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);
-val comp_eq_id_iff = result();
+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) \
@@ -426,16 +426,16 @@
 by (asm_full_simp_tac (ZF_ss 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));
-val fg_imp_bijective = result();
+qed "fg_imp_bijective";
 
 goal Perm.thy "!!f A. [| f: A->A;  f O f = id(A) |] ==> f : bij(A,A)";
 by (REPEAT (ares_tac [fg_imp_bijective] 1));
-val nilpotent_imp_bijective = result();
+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, 
 				  left_inverse_lemma, right_inverse_lemma]) 1);
-val invertible_imp_bijective = result();
+qed "invertible_imp_bijective";
 
 (** Unions of functions -- cf similar theorems on func.ML **)
 
@@ -444,7 +444,7 @@
 by (DEPTH_SOLVE_1 
       (resolve_tac [Un_least,converse_mono, Un_upper1,Un_upper2] 2));
 by (fast_tac ZF_cs 1);
-val converse_of_Un = result();
+qed "converse_of_Un";
 
 goalw Perm.thy [surj_def]
     "!!f g. [| f: surj(A,B);  g: surj(C,D);  A Int C = 0 |] ==> \
@@ -453,7 +453,7 @@
 	    ORELSE ball_tac 1
 	    ORELSE (rtac trans 1 THEN atac 2)
 	    ORELSE step_tac (ZF_cs addIs [fun_disjoint_Un]) 1));
-val surj_disjoint_Un = result();
+qed "surj_disjoint_Un";
 
 (*A simple, high-level proof; the version for injections follows from it,
   using  f:inj(A,B) <-> f:bij(A,range(f))  *)
@@ -463,7 +463,7 @@
 by (rtac invertible_imp_bijective 1);
 by (rtac (converse_of_Un RS ssubst) 1);
 by (REPEAT (ares_tac [fun_disjoint_Un, bij_is_fun, bij_converse_bij] 1));
-val bij_disjoint_Un = result();
+qed "bij_disjoint_Un";
 
 
 (** Restrictions as surjections and bijections *)
@@ -472,7 +472,7 @@
     "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);
-val surj_image = result();
+qed "surj_image";
 
 goal Perm.thy "!!f. [| f: Pi(C,B);  A<=C |] ==> restrict(f,A)``A = f``A";
 by (rtac equalityI 1);
@@ -480,21 +480,21 @@
 by (REPEAT (eresolve_tac [imageE, apply_equality RS subst] 2
      ORELSE ares_tac [subsetI,lamI,imageI] 2));
 by (REPEAT (ares_tac [image_mono,restrict_subset,subset_refl] 1));
-val restrict_image = result();
+qed "restrict_image";
 
 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 (REPEAT (eresolve_tac [asm_rl, bspec RS bspec RS mp, subsetD,
                           box_equals, restrict] 1));
-val restrict_inj = result();
+qed "restrict_inj";
 
 val prems = goal Perm.thy 
     "[| f: Pi(A,B);  C<=A |] ==> restrict(f,C): surj(C, f``C)";
 by (rtac (restrict_image RS subst) 1);
 by (rtac (restrict_type2 RS surj_image) 3);
 by (REPEAT (resolve_tac prems 1));
-val restrict_surj = result();
+qed "restrict_surj";
 
 goalw Perm.thy [inj_def,bij_def]
     "!!f. [| f: inj(A,B);  C<=A |] ==> restrict(f,C): bij(C, f``C)";
@@ -502,14 +502,14 @@
 by (REPEAT (eresolve_tac [bspec RS bspec RS mp, subsetD,
                           box_equals, restrict] 1
      ORELSE ares_tac [surj_is_fun,restrict_surj] 1));
-val restrict_bij = result();
+qed "restrict_bij";
 
 
 (*** 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);
-val inj_weaken_type = result();
+qed "inj_weaken_type";
 
 val [major] = goal Perm.thy  
     "[| f: inj(succ(m), A) |] ==> restrict(f,m) : inj(m, A-{f`m})";
@@ -524,7 +524,7 @@
 by (assume_tac 1);
 by (res_inst_tac [("a","m")] mem_irrefl 1);
 by (fast_tac ZF_cs 1);
-val inj_succ_restrict = result();
+qed "inj_succ_restrict";
 
 goalw Perm.thy [inj_def]
     "!!f. [| f: inj(A,B);  a~:A;  b~:B |]  ==> \
@@ -540,4 +540,4 @@
 by (ALLGOALS (asm_simp_tac 
 	      (FOL_ss addsimps [fun_extend_apply2,fun_extend_apply1])));
 by (ALLGOALS (fast_tac (ZF_cs addIs [apply_type])));
-val inj_extend = result();
+qed "inj_extend";