# HG changeset patch # User paulson # Date 1102590186 -3600 # Node ID aa785cea8fff0cd111c685a502af2e947274aa8f # Parent 24aff9e3de3f2beb27b2809bb1721d8498e65f12 converted Datatype_Universe to new-style theory diff -r 24aff9e3de3f -r aa785cea8fff src/HOL/Datatype_Universe.ML --- a/src/HOL/Datatype_Universe.ML Wed Dec 08 15:15:49 2004 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,561 +0,0 @@ -(* Title: HOL/Datatype_Universe.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1991 University of Cambridge -*) - -(** apfst -- can be used in similar type definitions **) - -Goalw [apfst_def] "apfst f (a,b) = (f(a),b)"; -by (rtac split_conv 1); -qed "apfst_conv"; - -val [major,minor] = Goal - "[| q = apfst f p; !!x y. [| p = (x,y); q = (f(x),y) |] ==> R \ -\ |] ==> R"; -by (rtac PairE 1); -by (rtac minor 1); -by (assume_tac 1); -by (rtac (major RS trans) 1); -by (etac ssubst 1); -by (rtac apfst_conv 1); -qed "apfst_convE"; - -(** Push -- an injection, analogous to Cons on lists **) - -Goalw [Push_def] "Push i f = Push j g ==> i=j"; -by (etac (fun_cong RS box_equals) 1); -by (rtac nat_case_0 1); -by (rtac nat_case_0 1); -qed "Push_inject1"; - -Goalw [Push_def] "Push i f = Push j g ==> f=g"; -by (rtac (ext RS box_equals) 1); -by (etac fun_cong 1); -by (rtac (nat_case_Suc RS ext) 1); -by (rtac (nat_case_Suc RS ext) 1); -qed "Push_inject2"; - -val [major,minor] = Goal - "[| Push i f =Push j g; [| i=j; f=g |] ==> P \ -\ |] ==> P"; -by (rtac ((major RS Push_inject2) RS ((major RS Push_inject1) RS minor)) 1); -qed "Push_inject"; - -Goalw [Push_def] "Push (Inr (Suc k)) f = (%z. Inr 0) ==> P"; -by (rtac Suc_neq_Zero 1); -by (etac (fun_cong RS box_equals RS Inr_inject) 1); -by (rtac nat_case_0 1); -by (rtac refl 1); -qed "Push_neq_K0"; - -(*** Isomorphisms ***) - -Goal "inj(Rep_Node)"; -by (rtac inj_inverseI 1); (*cannot combine by RS: multiple unifiers*) -by (rtac Rep_Node_inverse 1); -qed "inj_Rep_Node"; - -Goal "inj_on Abs_Node Node"; -by (rtac inj_on_inverseI 1); -by (etac Abs_Node_inverse 1); -qed "inj_on_Abs_Node"; - -bind_thm ("Abs_Node_inj", inj_on_Abs_Node RS inj_onD); - - -(*** Introduction rules for Node ***) - -Goalw [Node_def] "(%k. Inr 0, a) : Node"; -by (Blast_tac 1); -qed "Node_K0_I"; - -Goalw [Node_def,Push_def] - "p: Node ==> apfst (Push i) p : Node"; -by (fast_tac (claset() addSIs [apfst_conv, nat_case_Suc RS trans]) 1); -qed "Node_Push_I"; - - -(*** Distinctness of constructors ***) - -(** Scons vs Atom **) - -Goalw [Atom_def,Scons_def,Push_Node_def,One_nat_def] - "Scons M N ~= Atom(a)"; -by (rtac notI 1); -by (etac (equalityD2 RS subsetD RS UnE) 1); -by (rtac singletonI 1); -by (REPEAT (eresolve_tac [imageE, Abs_Node_inj RS apfst_convE, - Pair_inject, sym RS Push_neq_K0] 1 - ORELSE resolve_tac [Node_K0_I, Rep_Node RS Node_Push_I] 1)); -qed "Scons_not_Atom"; -bind_thm ("Atom_not_Scons", Scons_not_Atom RS not_sym); - - -(*** Injectiveness ***) - -(** Atomic nodes **) - -Goalw [Atom_def] "inj(Atom)"; -by (blast_tac (claset() addSIs [injI, Node_K0_I] addSDs [Abs_Node_inj]) 1); -qed "inj_Atom"; -bind_thm ("Atom_inject", inj_Atom RS injD); - -Goal "(Atom(a)=Atom(b)) = (a=b)"; -by (blast_tac (claset() addSDs [Atom_inject]) 1); -qed "Atom_Atom_eq"; -AddIffs [Atom_Atom_eq]; - -Goalw [Leaf_def,o_def] "inj(Leaf)"; -by (rtac injI 1); -by (etac (Atom_inject RS Inl_inject) 1); -qed "inj_Leaf"; - -bind_thm ("Leaf_inject", inj_Leaf RS injD); -AddSDs [Leaf_inject]; - -Goalw [Numb_def,o_def] "inj(Numb)"; -by (rtac injI 1); -by (etac (Atom_inject RS Inr_inject) 1); -qed "inj_Numb"; - -bind_thm ("Numb_inject", inj_Numb RS injD); -AddSDs [Numb_inject]; - -(** Injectiveness of Push_Node **) - -val [major,minor] = Goalw [Push_Node_def] - "[| Push_Node i m =Push_Node j n; [| i=j; m=n |] ==> P \ -\ |] ==> P"; -by (rtac (major RS Abs_Node_inj RS apfst_convE) 1); -by (REPEAT (resolve_tac [Rep_Node RS Node_Push_I] 1)); -by (etac (sym RS apfst_convE) 1); -by (rtac minor 1); -by (etac Pair_inject 1); -by (etac (Push_inject1 RS sym) 1); -by (rtac (inj_Rep_Node RS injD) 1); -by (etac trans 1); -by (safe_tac (claset() addSEs [Push_inject,sym])); -qed "Push_Node_inject"; - - -(** Injectiveness of Scons **) - -Goalw [Scons_def,One_nat_def] "Scons M N <= Scons M' N' ==> M<=M'"; -by (blast_tac (claset() addSDs [Push_Node_inject]) 1); -qed "Scons_inject_lemma1"; - -Goalw [Scons_def,One_nat_def] "Scons M N <= Scons M' N' ==> N<=N'"; -by (blast_tac (claset() addSDs [Push_Node_inject]) 1); -qed "Scons_inject_lemma2"; - -Goal "Scons M N = Scons M' N' ==> M=M'"; -by (etac equalityE 1); -by (REPEAT (ares_tac [equalityI, Scons_inject_lemma1] 1)); -qed "Scons_inject1"; - -Goal "Scons M N = Scons M' N' ==> N=N'"; -by (etac equalityE 1); -by (REPEAT (ares_tac [equalityI, Scons_inject_lemma2] 1)); -qed "Scons_inject2"; - -val [major,minor] = Goal - "[| Scons M N = Scons M' N'; [| M=M'; N=N' |] ==> P \ -\ |] ==> P"; -by (rtac ((major RS Scons_inject2) RS ((major RS Scons_inject1) RS minor)) 1); -qed "Scons_inject"; - -Goal "(Scons M N = Scons M' N') = (M=M' & N=N')"; -by (blast_tac (claset() addSEs [Scons_inject]) 1); -qed "Scons_Scons_eq"; - -(*** Distinctness involving Leaf and Numb ***) - -(** Scons vs Leaf **) - -Goalw [Leaf_def,o_def] "Scons M N ~= Leaf(a)"; -by (rtac Scons_not_Atom 1); -qed "Scons_not_Leaf"; -bind_thm ("Leaf_not_Scons", Scons_not_Leaf RS not_sym); - -AddIffs [Scons_not_Leaf, Leaf_not_Scons]; - - -(** Scons vs Numb **) - -Goalw [Numb_def,o_def] "Scons M N ~= Numb(k)"; -by (rtac Scons_not_Atom 1); -qed "Scons_not_Numb"; -bind_thm ("Numb_not_Scons", Scons_not_Numb RS not_sym); - -AddIffs [Scons_not_Numb, Numb_not_Scons]; - - -(** Leaf vs Numb **) - -Goalw [Leaf_def,Numb_def] "Leaf(a) ~= Numb(k)"; -by (simp_tac (simpset() addsimps [Inl_not_Inr]) 1); -qed "Leaf_not_Numb"; -bind_thm ("Numb_not_Leaf", Leaf_not_Numb RS not_sym); - -AddIffs [Leaf_not_Numb, Numb_not_Leaf]; - - -(*** ndepth -- the depth of a node ***) - -Addsimps [apfst_conv]; -AddIffs [Scons_not_Atom, Atom_not_Scons, Scons_Scons_eq]; - - -Goalw [ndepth_def] "ndepth (Abs_Node(%k. Inr 0, x)) = 0"; -by (EVERY1[stac (Node_K0_I RS Abs_Node_inverse), stac split_conv]); -by (rtac Least_equality 1); -by Auto_tac; -qed "ndepth_K0"; - -Goal "nat_case (Inr (Suc i)) f k = Inr 0 --> Suc(LEAST x. f x = Inr 0) <= k"; -by (induct_tac "k" 1); -by (ALLGOALS Simp_tac); -by (rtac impI 1); -by (etac Least_le 1); -val lemma = result(); - -Goalw [ndepth_def,Push_Node_def] - "ndepth (Push_Node (Inr (Suc i)) n) = Suc(ndepth(n))"; -by (stac (Rep_Node RS Node_Push_I RS Abs_Node_inverse) 1); -by (cut_facts_tac [rewrite_rule [Node_def] Rep_Node] 1); -by Safe_tac; -by (etac ssubst 1); (*instantiates type variables!*) -by (Simp_tac 1); -by (rtac Least_equality 1); -by (rewtac Push_def); -by (auto_tac (claset(), simpset() addsimps [lemma])); -by (etac LeastI 1); -qed "ndepth_Push_Node"; - - -(*** ntrunc applied to the various node sets ***) - -Goalw [ntrunc_def] "ntrunc 0 M = {}"; -by (Blast_tac 1); -qed "ntrunc_0"; - -Goalw [Atom_def,ntrunc_def] "ntrunc (Suc k) (Atom a) = Atom(a)"; -by (fast_tac (claset() addss (simpset() addsimps [ndepth_K0])) 1); -qed "ntrunc_Atom"; - -Goalw [Leaf_def,o_def] "ntrunc (Suc k) (Leaf a) = Leaf(a)"; -by (rtac ntrunc_Atom 1); -qed "ntrunc_Leaf"; - -Goalw [Numb_def,o_def] "ntrunc (Suc k) (Numb i) = Numb(i)"; -by (rtac ntrunc_Atom 1); -qed "ntrunc_Numb"; - -Goalw [Scons_def,ntrunc_def,One_nat_def] - "ntrunc (Suc k) (Scons M N) = Scons (ntrunc k M) (ntrunc k N)"; -by (safe_tac (claset() addSIs [imageI])); -by (REPEAT (stac ndepth_Push_Node 3 THEN etac Suc_mono 3)); -by (REPEAT (rtac Suc_less_SucD 1 THEN - rtac (ndepth_Push_Node RS subst) 1 THEN - assume_tac 1)); -qed "ntrunc_Scons"; - -Addsimps [ntrunc_0, ntrunc_Atom, ntrunc_Leaf, ntrunc_Numb, ntrunc_Scons]; - - -(** Injection nodes **) - -Goalw [In0_def] "ntrunc (Suc 0) (In0 M) = {}"; -by (Simp_tac 1); -by (rewtac Scons_def); -by (Blast_tac 1); -qed "ntrunc_one_In0"; - -Goalw [In0_def] - "ntrunc (Suc (Suc k)) (In0 M) = In0 (ntrunc (Suc k) M)"; -by (Simp_tac 1); -qed "ntrunc_In0"; - -Goalw [In1_def] "ntrunc (Suc 0) (In1 M) = {}"; -by (Simp_tac 1); -by (rewtac Scons_def); -by (Blast_tac 1); -qed "ntrunc_one_In1"; - -Goalw [In1_def] - "ntrunc (Suc (Suc k)) (In1 M) = In1 (ntrunc (Suc k) M)"; -by (Simp_tac 1); -qed "ntrunc_In1"; - -Addsimps [ntrunc_one_In0, ntrunc_In0, ntrunc_one_In1, ntrunc_In1]; - - -(*** Cartesian Product ***) - -Goalw [uprod_def] "[| M:A; N:B |] ==> Scons M N : uprod A B"; -by (REPEAT (ares_tac [singletonI,UN_I] 1)); -qed "uprodI"; - -(*The general elimination rule*) -val major::prems = Goalw [uprod_def] - "[| c : uprod A B; \ -\ !!x y. [| x:A; y:B; c = Scons x y |] ==> P \ -\ |] ==> P"; -by (cut_facts_tac [major] 1); -by (REPEAT (eresolve_tac [asm_rl,singletonE,UN_E] 1 - ORELSE resolve_tac prems 1)); -qed "uprodE"; - -(*Elimination of a pair -- introduces no eigenvariables*) -val prems = Goal - "[| Scons M N : uprod A B; [| M:A; N:B |] ==> P \ -\ |] ==> P"; -by (rtac uprodE 1); -by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Scons_inject,ssubst] 1)); -qed "uprodE2"; - - -(*** Disjoint Sum ***) - -Goalw [usum_def] "M:A ==> In0(M) : usum A B"; -by (Blast_tac 1); -qed "usum_In0I"; - -Goalw [usum_def] "N:B ==> In1(N) : usum A B"; -by (Blast_tac 1); -qed "usum_In1I"; - -val major::prems = Goalw [usum_def] - "[| u : usum A B; \ -\ !!x. [| x:A; u=In0(x) |] ==> P; \ -\ !!y. [| y:B; u=In1(y) |] ==> P \ -\ |] ==> P"; -by (rtac (major RS UnE) 1); -by (REPEAT (rtac refl 1 - ORELSE eresolve_tac (prems@[imageE,ssubst]) 1)); -qed "usumE"; - - -(** Injection **) - -Goalw [In0_def,In1_def,One_nat_def] "In0(M) ~= In1(N)"; -by (rtac notI 1); -by (etac (Scons_inject1 RS Numb_inject RS Zero_neq_Suc) 1); -qed "In0_not_In1"; - -bind_thm ("In1_not_In0", In0_not_In1 RS not_sym); - -AddIffs [In0_not_In1, In1_not_In0]; - -Goalw [In0_def] "In0(M) = In0(N) ==> M=N"; -by (etac (Scons_inject2) 1); -qed "In0_inject"; - -Goalw [In1_def] "In1(M) = In1(N) ==> M=N"; -by (etac (Scons_inject2) 1); -qed "In1_inject"; - -Goal "(In0 M = In0 N) = (M=N)"; -by (blast_tac (claset() addSDs [In0_inject]) 1); -qed "In0_eq"; - -Goal "(In1 M = In1 N) = (M=N)"; -by (blast_tac (claset() addSDs [In1_inject]) 1); -qed "In1_eq"; - -AddIffs [In0_eq, In1_eq]; - -Goal "inj In0"; -by (blast_tac (claset() addSIs [injI]) 1); -qed "inj_In0"; - -Goal "inj In1"; -by (blast_tac (claset() addSIs [injI]) 1); -qed "inj_In1"; - - -(*** Function spaces ***) - -Goalw [Lim_def] "Lim f = Lim g ==> f = g"; -by (rtac ext 1); -by (blast_tac (claset() addSEs [Push_Node_inject]) 1); -qed "Lim_inject"; - - -(*** proving equality of sets and functions using ntrunc ***) - -Goalw [ntrunc_def] "ntrunc k M <= M"; -by (Blast_tac 1); -qed "ntrunc_subsetI"; - -val [major] = Goalw [ntrunc_def] "(!!k. ntrunc k M <= N) ==> M<=N"; -by (blast_tac (claset() addIs [less_add_Suc1, less_add_Suc2, - major RS subsetD]) 1); -qed "ntrunc_subsetD"; - -(*A generalized form of the take-lemma*) -val [major] = Goal "(!!k. ntrunc k M = ntrunc k N) ==> M=N"; -by (rtac equalityI 1); -by (ALLGOALS (rtac ntrunc_subsetD)); -by (ALLGOALS (rtac (ntrunc_subsetI RSN (2, subset_trans)))); -by (rtac (major RS equalityD1) 1); -by (rtac (major RS equalityD2) 1); -qed "ntrunc_equality"; - -val [major] = Goalw [o_def] - "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2"; -by (rtac (ntrunc_equality RS ext) 1); -by (rtac (major RS fun_cong) 1); -qed "ntrunc_o_equality"; - -(*** Monotonicity ***) - -Goalw [uprod_def] "[| A<=A'; B<=B' |] ==> uprod A B <= uprod A' B'"; -by (Blast_tac 1); -qed "uprod_mono"; - -Goalw [usum_def] "[| A<=A'; B<=B' |] ==> usum A B <= usum A' B'"; -by (Blast_tac 1); -qed "usum_mono"; - -Goalw [Scons_def] "[| M<=M'; N<=N' |] ==> Scons M N <= Scons M' N'"; -by (Blast_tac 1); -qed "Scons_mono"; - -Goalw [In0_def] "M<=N ==> In0(M) <= In0(N)"; -by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); -qed "In0_mono"; - -Goalw [In1_def] "M<=N ==> In1(M) <= In1(N)"; -by (REPEAT (ares_tac [subset_refl,Scons_mono] 1)); -qed "In1_mono"; - - -(*** Split and Case ***) - -Goalw [Split_def] "Split c (Scons M N) = c M N"; -by (Blast_tac 1); -qed "Split"; - -Goalw [Case_def] "Case c d (In0 M) = c(M)"; -by (Blast_tac 1); -qed "Case_In0"; - -Goalw [Case_def] "Case c d (In1 N) = d(N)"; -by (Blast_tac 1); -qed "Case_In1"; - -Addsimps [Split, Case_In0, Case_In1]; - - -(**** UN x. B(x) rules ****) - -Goalw [ntrunc_def] "ntrunc k (UN x. f(x)) = (UN x. ntrunc k (f x))"; -by (Blast_tac 1); -qed "ntrunc_UN1"; - -Goalw [Scons_def] "Scons (UN x. f x) M = (UN x. Scons (f x) M)"; -by (Blast_tac 1); -qed "Scons_UN1_x"; - -Goalw [Scons_def] "Scons M (UN x. f x) = (UN x. Scons M (f x))"; -by (Blast_tac 1); -qed "Scons_UN1_y"; - -Goalw [In0_def] "In0(UN x. f(x)) = (UN x. In0(f(x)))"; -by (rtac Scons_UN1_y 1); -qed "In0_UN1"; - -Goalw [In1_def] "In1(UN x. f(x)) = (UN x. In1(f(x)))"; -by (rtac Scons_UN1_y 1); -qed "In1_UN1"; - - -(*** Equality for Cartesian Product ***) - -Goalw [dprod_def] - "[| (M,M'):r; (N,N'):s |] ==> (Scons M N, Scons M' N') : dprod r s"; -by (Blast_tac 1); -qed "dprodI"; - -(*The general elimination rule*) -val major::prems = Goalw [dprod_def] - "[| c : dprod r s; \ -\ !!x y x' y'. [| (x,x') : r; (y,y') : s; c = (Scons x y, Scons x' y') |] ==> P \ -\ |] ==> P"; -by (cut_facts_tac [major] 1); -by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, mem_splitE, singletonE])); -by (REPEAT (ares_tac prems 1 ORELSE hyp_subst_tac 1)); -qed "dprodE"; - - -(*** Equality for Disjoint Sum ***) - -Goalw [dsum_def] "(M,M'):r ==> (In0(M), In0(M')) : dsum r s"; -by (Blast_tac 1); -qed "dsum_In0I"; - -Goalw [dsum_def] "(N,N'):s ==> (In1(N), In1(N')) : dsum r s"; -by (Blast_tac 1); -qed "dsum_In1I"; - -val major::prems = Goalw [dsum_def] - "[| w : dsum r s; \ -\ !!x x'. [| (x,x') : r; w = (In0(x), In0(x')) |] ==> P; \ -\ !!y y'. [| (y,y') : s; w = (In1(y), In1(y')) |] ==> P \ -\ |] ==> P"; -by (cut_facts_tac [major] 1); -by (REPEAT_FIRST (eresolve_tac [asm_rl, UN_E, UnE, mem_splitE, singletonE])); -by (DEPTH_SOLVE (ares_tac prems 1 ORELSE hyp_subst_tac 1)); -qed "dsumE"; - -AddSIs [uprodI, dprodI]; -AddIs [usum_In0I, usum_In1I, dsum_In0I, dsum_In1I]; -AddSEs [uprodE, dprodE, usumE, dsumE]; - - -(*** Monotonicity ***) - -Goal "[| r<=r'; s<=s' |] ==> dprod r s <= dprod r' s'"; -by (Blast_tac 1); -qed "dprod_mono"; - -Goal "[| r<=r'; s<=s' |] ==> dsum r s <= dsum r' s'"; -by (Blast_tac 1); -qed "dsum_mono"; - - -(*** Bounding theorems ***) - -Goal "(dprod (A <*> B) (C <*> D)) <= (uprod A C) <*> (uprod B D)"; -by (Blast_tac 1); -qed "dprod_Sigma"; - -bind_thm ("dprod_subset_Sigma", [dprod_mono, dprod_Sigma] MRS subset_trans |> standard); - -(*Dependent version*) -Goal "(dprod (Sigma A B) (Sigma C D)) <= Sigma (uprod A C) (Split (%x y. uprod (B x) (D y)))"; -by Safe_tac; -by (stac Split 1); -by (Blast_tac 1); -qed "dprod_subset_Sigma2"; - -Goal "(dsum (A <*> B) (C <*> D)) <= (usum A C) <*> (usum B D)"; -by (Blast_tac 1); -qed "dsum_Sigma"; - -bind_thm ("dsum_subset_Sigma", [dsum_mono, dsum_Sigma] MRS subset_trans |> standard); - - -(*** Domain ***) - -Goal "Domain (dprod r s) = uprod (Domain r) (Domain s)"; -by Auto_tac; -qed "Domain_dprod"; - -Goal "Domain (dsum r s) = usum (Domain r) (Domain s)"; -by Auto_tac; -qed "Domain_dsum"; - -Addsimps [Domain_dprod, Domain_dsum]; diff -r 24aff9e3de3f -r aa785cea8fff src/HOL/Datatype_Universe.thy --- a/src/HOL/Datatype_Universe.thy Wed Dec 08 15:15:49 2004 +0100 +++ b/src/HOL/Datatype_Universe.thy Thu Dec 09 12:03:06 2004 +0100 @@ -3,97 +3,637 @@ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge -Declares the type ('a, 'b) node, a subtype of (nat=>'b+nat) * ('a+nat) - -Defines "Cartesian Product" and "Disjoint Sum" as set operations. Could <*> be generalized to a general summation (Sigma)? *) -Datatype_Universe = NatArith + Sum_Type + +header{*Analogues of the Cartesian Product and Disjoint Sum for Datatypes*} +theory Datatype_Universe +imports NatArith Sum_Type +begin -(** lists, trees will be sets of nodes **) typedef (Node) ('a,'b) node = "{p. EX f x k. p = (f::nat=>'b+nat, x::'a+nat) & f k = Inr 0}" + --{*it is a subtype of @{text "(nat=>'b+nat) * ('a+nat)"}*} + by auto -types - 'a item = ('a, unit) node set - ('a, 'b) dtree = ('a, 'b) node set +text{*Datatypes will be represented by sets of type @{text node}*} + +types 'a item = "('a, unit) node set" + ('a, 'b) dtree = "('a, 'b) node set" consts apfst :: "['a=>'c, 'a*'b] => 'c*'b" Push :: "[('b + nat), nat => ('b + nat)] => (nat => ('b + nat))" Push_Node :: "[('b + nat), ('a, 'b) node] => ('a, 'b) node" - ndepth :: ('a, 'b) node => nat + ndepth :: "('a, 'b) node => nat" Atom :: "('a + nat) => ('a, 'b) dtree" - Leaf :: 'a => ('a, 'b) dtree - Numb :: nat => ('a, 'b) dtree - Scons :: [('a, 'b) dtree, ('a, 'b) dtree] => ('a, 'b) dtree - In0,In1 :: ('a, 'b) dtree => ('a, 'b) dtree - Lim :: ('b => ('a, 'b) dtree) => ('a, 'b) dtree + Leaf :: "'a => ('a, 'b) dtree" + Numb :: "nat => ('a, 'b) dtree" + Scons :: "[('a, 'b) dtree, ('a, 'b) dtree] => ('a, 'b) dtree" + In0 :: "('a, 'b) dtree => ('a, 'b) dtree" + In1 :: "('a, 'b) dtree => ('a, 'b) dtree" + Lim :: "('b => ('a, 'b) dtree) => ('a, 'b) dtree" - ntrunc :: [nat, ('a, 'b) dtree] => ('a, 'b) dtree + ntrunc :: "[nat, ('a, 'b) dtree] => ('a, 'b) dtree" - uprod :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set - usum :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set + uprod :: "[('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set" + usum :: "[('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set" - Split :: [[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c - Case :: [[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c + Split :: "[[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c" + Case :: "[[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c" - dprod :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] + dprod :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] => (('a, 'b) dtree * ('a, 'b) dtree)set" - dsum :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] + dsum :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] => (('a, 'b) dtree * ('a, 'b) dtree)set" defs - Push_Node_def "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))" + Push_Node_def: "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))" (*crude "lists" of nats -- needed for the constructions*) - apfst_def "apfst == (%f (x,y). (f(x),y))" - Push_def "Push == (%b h. nat_case b h)" + apfst_def: "apfst == (%f (x,y). (f(x),y))" + Push_def: "Push == (%b h. nat_case b h)" (** operations on S-expressions -- sets of nodes **) (*S-expression constructors*) - Atom_def "Atom == (%x. {Abs_Node((%k. Inr 0, x))})" - Scons_def "Scons M N == (Push_Node (Inr 1) ` M) Un (Push_Node (Inr (Suc 1)) ` N)" + Atom_def: "Atom == (%x. {Abs_Node((%k. Inr 0, x))})" + Scons_def: "Scons M N == (Push_Node (Inr 1) ` M) Un (Push_Node (Inr (Suc 1)) ` N)" (*Leaf nodes, with arbitrary or nat labels*) - Leaf_def "Leaf == Atom o Inl" - Numb_def "Numb == Atom o Inr" + Leaf_def: "Leaf == Atom o Inl" + Numb_def: "Numb == Atom o Inr" (*Injections of the "disjoint sum"*) - In0_def "In0(M) == Scons (Numb 0) M" - In1_def "In1(M) == Scons (Numb 1) M" + In0_def: "In0(M) == Scons (Numb 0) M" + In1_def: "In1(M) == Scons (Numb 1) M" (*Function spaces*) - Lim_def "Lim f == Union {z. ? x. z = Push_Node (Inl x) ` (f x)}" + Lim_def: "Lim f == Union {z. ? x. z = Push_Node (Inl x) ` (f x)}" (*the set of nodes with depth less than k*) - ndepth_def "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)" - ntrunc_def "ntrunc k N == {n. n:N & ndepth(n) R + |] ==> R" +by (force simp add: apfst_def) + +(** Push -- an injection, analogous to Cons on lists **) + +lemma Push_inject1: "Push i f = Push j g ==> i=j" +apply (simp add: Push_def expand_fun_eq) +apply (drule_tac x=0 in spec, simp) +done + +lemma Push_inject2: "Push i f = Push j g ==> f=g" +apply (auto simp add: Push_def expand_fun_eq) +apply (drule_tac x="Suc x" in spec, simp) +done + +lemma Push_inject: + "[| Push i f =Push j g; [| i=j; f=g |] ==> P |] ==> P" +by (blast dest: Push_inject1 Push_inject2) + +lemma Push_neq_K0: "Push (Inr (Suc k)) f = (%z. Inr 0) ==> P" +by (auto simp add: Push_def expand_fun_eq split: nat.split_asm) + +(*** Isomorphisms ***) + +lemma inj_Rep_Node: "inj(Rep_Node)" +apply (rule inj_on_inverseI) +apply (rule Rep_Node_inverse) +done + +lemma inj_on_Abs_Node: "inj_on Abs_Node Node" +apply (rule inj_on_inverseI) +apply (erule Abs_Node_inverse) +done + +lemmas Abs_Node_inj = inj_on_Abs_Node [THEN inj_onD, standard] + + +(*** Introduction rules for Node ***) + +lemma Node_K0_I: "(%k. Inr 0, a) : Node" +by (simp add: Node_def) + +lemma Node_Push_I: + "p: Node ==> apfst (Push i) p : Node" +apply (simp add: Node_def Push_def) +apply (fast intro!: apfst_conv nat_case_Suc [THEN trans]) +done + + +subsubsection{*Freeness: Distinctness of Constructors*} + +(** Scons vs Atom **) + +lemma Scons_not_Atom [iff]: "Scons M N \ Atom(a)" +apply (simp add: Atom_def Scons_def Push_Node_def One_nat_def) +apply (blast intro: Node_K0_I Rep_Node [THEN Node_Push_I] + dest!: Abs_Node_inj + elim!: apfst_convE sym [THEN Push_neq_K0]) +done + +lemmas Atom_not_Scons = Scons_not_Atom [THEN not_sym, standard, iff] + + +(*** Injectiveness ***) + +(** Atomic nodes **) + +lemma inj_Atom: "inj(Atom)" +apply (simp add: Atom_def) +apply (blast intro!: inj_onI Node_K0_I dest!: Abs_Node_inj) +done +lemmas Atom_inject = inj_Atom [THEN injD, standard] + +lemma Atom_Atom_eq [iff]: "(Atom(a)=Atom(b)) = (a=b)" +by (blast dest!: Atom_inject) + +lemma inj_Leaf: "inj(Leaf)" +apply (simp add: Leaf_def o_def) +apply (rule inj_onI) +apply (erule Atom_inject [THEN Inl_inject]) +done + +lemmas Leaf_inject = inj_Leaf [THEN injD, standard, dest!] + +lemma inj_Numb: "inj(Numb)" +apply (simp add: Numb_def o_def) +apply (rule inj_onI) +apply (erule Atom_inject [THEN Inr_inject]) +done + +lemmas Numb_inject = inj_Numb [THEN injD, standard, dest!] + + +(** Injectiveness of Push_Node **) + +lemma Push_Node_inject: + "[| Push_Node i m =Push_Node j n; [| i=j; m=n |] ==> P + |] ==> P" +apply (simp add: Push_Node_def) +apply (erule Abs_Node_inj [THEN apfst_convE]) +apply (rule Rep_Node [THEN Node_Push_I])+ +apply (erule sym [THEN apfst_convE]) +apply (blast intro: inj_Rep_Node [THEN injD] trans sym elim!: Push_inject) +done + + +(** Injectiveness of Scons **) + +lemma Scons_inject_lemma1: "Scons M N <= Scons M' N' ==> M<=M'" +apply (simp add: Scons_def One_nat_def) +apply (blast dest!: Push_Node_inject) +done + +lemma Scons_inject_lemma2: "Scons M N <= Scons M' N' ==> N<=N'" +apply (simp add: Scons_def One_nat_def) +apply (blast dest!: Push_Node_inject) +done + +lemma Scons_inject1: "Scons M N = Scons M' N' ==> M=M'" +apply (erule equalityE) +apply (rules intro: equalityI Scons_inject_lemma1) +done + +lemma Scons_inject2: "Scons M N = Scons M' N' ==> N=N'" +apply (erule equalityE) +apply (rules intro: equalityI Scons_inject_lemma2) +done + +lemma Scons_inject: + "[| Scons M N = Scons M' N'; [| M=M'; N=N' |] ==> P |] ==> P" +by (rules dest: Scons_inject1 Scons_inject2) + +lemma Scons_Scons_eq [iff]: "(Scons M N = Scons M' N') = (M=M' & N=N')" +by (blast elim!: Scons_inject) + +(*** Distinctness involving Leaf and Numb ***) + +(** Scons vs Leaf **) + +lemma Scons_not_Leaf [iff]: "Scons M N \ Leaf(a)" +by (simp add: Leaf_def o_def Scons_not_Atom) + +lemmas Leaf_not_Scons = Scons_not_Leaf [THEN not_sym, standard, iff] + + +(** Scons vs Numb **) + +lemma Scons_not_Numb [iff]: "Scons M N \ Numb(k)" +by (simp add: Numb_def o_def Scons_not_Atom) + +lemmas Numb_not_Scons = Scons_not_Numb [THEN not_sym, standard, iff] + + + +(** Leaf vs Numb **) + +lemma Leaf_not_Numb [iff]: "Leaf(a) \ Numb(k)" +by (simp add: Leaf_def Numb_def) + +lemmas Numb_not_Leaf = Leaf_not_Numb [THEN not_sym, standard, iff] + + + +(*** ndepth -- the depth of a node ***) + +lemma ndepth_K0: "ndepth (Abs_Node(%k. Inr 0, x)) = 0" +by (simp add: ndepth_def Node_K0_I [THEN Abs_Node_inverse] Least_equality) + +lemma ndepth_Push_Node_aux: + "nat_case (Inr (Suc i)) f k = Inr 0 --> Suc(LEAST x. f x = Inr 0) <= k" +apply (induct_tac "k", auto) +apply (erule Least_le) +done + +lemma ndepth_Push_Node: + "ndepth (Push_Node (Inr (Suc i)) n) = Suc(ndepth(n))" +apply (insert Rep_Node [of n, unfolded Node_def]) +apply (auto simp add: ndepth_def Push_Node_def + Rep_Node [THEN Node_Push_I, THEN Abs_Node_inverse]) +apply (rule Least_equality) +apply (auto simp add: Push_def ndepth_Push_Node_aux) +apply (erule LeastI) +done + + +(*** ntrunc applied to the various node sets ***) + +lemma ntrunc_0 [simp]: "ntrunc 0 M = {}" +by (simp add: ntrunc_def) + +lemma ntrunc_Atom [simp]: "ntrunc (Suc k) (Atom a) = Atom(a)" +by (auto simp add: Atom_def ntrunc_def ndepth_K0) + +lemma ntrunc_Leaf [simp]: "ntrunc (Suc k) (Leaf a) = Leaf(a)" +by (simp add: Leaf_def o_def ntrunc_Atom) + +lemma ntrunc_Numb [simp]: "ntrunc (Suc k) (Numb i) = Numb(i)" +by (simp add: Numb_def o_def ntrunc_Atom) + +lemma ntrunc_Scons [simp]: + "ntrunc (Suc k) (Scons M N) = Scons (ntrunc k M) (ntrunc k N)" +by (auto simp add: Scons_def ntrunc_def One_nat_def ndepth_Push_Node) + + + +(** Injection nodes **) + +lemma ntrunc_one_In0 [simp]: "ntrunc (Suc 0) (In0 M) = {}" +apply (simp add: In0_def) +apply (simp add: Scons_def) +done + +lemma ntrunc_In0 [simp]: "ntrunc (Suc(Suc k)) (In0 M) = In0 (ntrunc (Suc k) M)" +by (simp add: In0_def) + +lemma ntrunc_one_In1 [simp]: "ntrunc (Suc 0) (In1 M) = {}" +apply (simp add: In1_def) +apply (simp add: Scons_def) +done + +lemma ntrunc_In1 [simp]: "ntrunc (Suc(Suc k)) (In1 M) = In1 (ntrunc (Suc k) M)" +by (simp add: In1_def) + + +subsection{*Set Constructions*} + + +(*** Cartesian Product ***) + +lemma uprodI [intro!]: "[| M:A; N:B |] ==> Scons M N : uprod A B" +by (simp add: uprod_def) + +(*The general elimination rule*) +lemma uprodE [elim!]: + "[| c : uprod A B; + !!x y. [| x:A; y:B; c = Scons x y |] ==> P + |] ==> P" +by (auto simp add: uprod_def) + + +(*Elimination of a pair -- introduces no eigenvariables*) +lemma uprodE2: "[| Scons M N : uprod A B; [| M:A; N:B |] ==> P |] ==> P" +by (auto simp add: uprod_def) + + +(*** Disjoint Sum ***) + +lemma usum_In0I [intro]: "M:A ==> In0(M) : usum A B" +by (simp add: usum_def) + +lemma usum_In1I [intro]: "N:B ==> In1(N) : usum A B" +by (simp add: usum_def) + +lemma usumE [elim!]: + "[| u : usum A B; + !!x. [| x:A; u=In0(x) |] ==> P; + !!y. [| y:B; u=In1(y) |] ==> P + |] ==> P" +by (auto simp add: usum_def) + + +(** Injection **) + +lemma In0_not_In1 [iff]: "In0(M) \ In1(N)" +by (auto simp add: In0_def In1_def One_nat_def) + +lemmas In1_not_In0 = In0_not_In1 [THEN not_sym, standard, iff] + +lemma In0_inject: "In0(M) = In0(N) ==> M=N" +by (simp add: In0_def) + +lemma In1_inject: "In1(M) = In1(N) ==> M=N" +by (simp add: In1_def) + +lemma In0_eq [iff]: "(In0 M = In0 N) = (M=N)" +by (blast dest!: In0_inject) + +lemma In1_eq [iff]: "(In1 M = In1 N) = (M=N)" +by (blast dest!: In1_inject) + +lemma inj_In0: "inj In0" +by (blast intro!: inj_onI) + +lemma inj_In1: "inj In1" +by (blast intro!: inj_onI) + + +(*** Function spaces ***) + +lemma Lim_inject: "Lim f = Lim g ==> f = g" +apply (simp add: Lim_def) +apply (rule ext) +apply (blast elim!: Push_Node_inject) +done + + +(*** proving equality of sets and functions using ntrunc ***) + +lemma ntrunc_subsetI: "ntrunc k M <= M" +by (auto simp add: ntrunc_def) + +lemma ntrunc_subsetD: "(!!k. ntrunc k M <= N) ==> M<=N" +by (auto simp add: ntrunc_def) + +(*A generalized form of the take-lemma*) +lemma ntrunc_equality: "(!!k. ntrunc k M = ntrunc k N) ==> M=N" +apply (rule equalityI) +apply (rule_tac [!] ntrunc_subsetD) +apply (rule_tac [!] ntrunc_subsetI [THEN [2] subset_trans], auto) +done + +lemma ntrunc_o_equality: + "[| !!k. (ntrunc(k) o h1) = (ntrunc(k) o h2) |] ==> h1=h2" +apply (rule ntrunc_equality [THEN ext]) +apply (simp add: expand_fun_eq) +done + + +(*** Monotonicity ***) + +lemma uprod_mono: "[| A<=A'; B<=B' |] ==> uprod A B <= uprod A' B'" +by (simp add: uprod_def, blast) + +lemma usum_mono: "[| A<=A'; B<=B' |] ==> usum A B <= usum A' B'" +by (simp add: usum_def, blast) + +lemma Scons_mono: "[| M<=M'; N<=N' |] ==> Scons M N <= Scons M' N'" +by (simp add: Scons_def, blast) + +lemma In0_mono: "M<=N ==> In0(M) <= In0(N)" +by (simp add: In0_def subset_refl Scons_mono) + +lemma In1_mono: "M<=N ==> In1(M) <= In1(N)" +by (simp add: In1_def subset_refl Scons_mono) + + +(*** Split and Case ***) + +lemma Split [simp]: "Split c (Scons M N) = c M N" +by (simp add: Split_def) + +lemma Case_In0 [simp]: "Case c d (In0 M) = c(M)" +by (simp add: Case_def) + +lemma Case_In1 [simp]: "Case c d (In1 N) = d(N)" +by (simp add: Case_def) + + + +(**** UN x. B(x) rules ****) + +lemma ntrunc_UN1: "ntrunc k (UN x. f(x)) = (UN x. ntrunc k (f x))" +by (simp add: ntrunc_def, blast) + +lemma Scons_UN1_x: "Scons (UN x. f x) M = (UN x. Scons (f x) M)" +by (simp add: Scons_def, blast) + +lemma Scons_UN1_y: "Scons M (UN x. f x) = (UN x. Scons M (f x))" +by (simp add: Scons_def, blast) + +lemma In0_UN1: "In0(UN x. f(x)) = (UN x. In0(f(x)))" +by (simp add: In0_def Scons_UN1_y) + +lemma In1_UN1: "In1(UN x. f(x)) = (UN x. In1(f(x)))" +by (simp add: In1_def Scons_UN1_y) + + +(*** Equality for Cartesian Product ***) + +lemma dprodI [intro!]: + "[| (M,M'):r; (N,N'):s |] ==> (Scons M N, Scons M' N') : dprod r s" +by (auto simp add: dprod_def) + +(*The general elimination rule*) +lemma dprodE [elim!]: + "[| c : dprod r s; + !!x y x' y'. [| (x,x') : r; (y,y') : s; + c = (Scons x y, Scons x' y') |] ==> P + |] ==> P" +by (auto simp add: dprod_def) + + +(*** Equality for Disjoint Sum ***) + +lemma dsum_In0I [intro]: "(M,M'):r ==> (In0(M), In0(M')) : dsum r s" +by (auto simp add: dsum_def) + +lemma dsum_In1I [intro]: "(N,N'):s ==> (In1(N), In1(N')) : dsum r s" +by (auto simp add: dsum_def) + +lemma dsumE [elim!]: + "[| w : dsum r s; + !!x x'. [| (x,x') : r; w = (In0(x), In0(x')) |] ==> P; + !!y y'. [| (y,y') : s; w = (In1(y), In1(y')) |] ==> P + |] ==> P" +by (auto simp add: dsum_def) + + +(*** Monotonicity ***) + +lemma dprod_mono: "[| r<=r'; s<=s' |] ==> dprod r s <= dprod r' s'" +by blast + +lemma dsum_mono: "[| r<=r'; s<=s' |] ==> dsum r s <= dsum r' s'" +by blast + + +(*** Bounding theorems ***) + +lemma dprod_Sigma: "(dprod (A <*> B) (C <*> D)) <= (uprod A C) <*> (uprod B D)" +by blast + +lemmas dprod_subset_Sigma = subset_trans [OF dprod_mono dprod_Sigma, standard] + +(*Dependent version*) +lemma dprod_subset_Sigma2: + "(dprod (Sigma A B) (Sigma C D)) <= + Sigma (uprod A C) (Split (%x y. uprod (B x) (D y)))" +by auto + +lemma dsum_Sigma: "(dsum (A <*> B) (C <*> D)) <= (usum A C) <*> (usum B D)" +by blast + +lemmas dsum_subset_Sigma = subset_trans [OF dsum_mono dsum_Sigma, standard] + + +(*** Domain ***) + +lemma Domain_dprod [simp]: "Domain (dprod r s) = uprod (Domain r) (Domain s)" +by auto + +lemma Domain_dsum [simp]: "Domain (dsum r s) = usum (Domain r) (Domain s)" +by auto + +ML +{* +val apfst_conv = thm "apfst_conv"; +val apfst_convE = thm "apfst_convE"; +val Push_inject1 = thm "Push_inject1"; +val Push_inject2 = thm "Push_inject2"; +val Push_inject = thm "Push_inject"; +val Push_neq_K0 = thm "Push_neq_K0"; +val inj_Rep_Node = thm "inj_Rep_Node"; +val inj_on_Abs_Node = thm "inj_on_Abs_Node"; +val Abs_Node_inj = thm "Abs_Node_inj"; +val Node_K0_I = thm "Node_K0_I"; +val Node_Push_I = thm "Node_Push_I"; +val Scons_not_Atom = thm "Scons_not_Atom"; +val Atom_not_Scons = thm "Atom_not_Scons"; +val inj_Atom = thm "inj_Atom"; +val Atom_inject = thm "Atom_inject"; +val Atom_Atom_eq = thm "Atom_Atom_eq"; +val inj_Leaf = thm "inj_Leaf"; +val Leaf_inject = thm "Leaf_inject"; +val inj_Numb = thm "inj_Numb"; +val Numb_inject = thm "Numb_inject"; +val Push_Node_inject = thm "Push_Node_inject"; +val Scons_inject1 = thm "Scons_inject1"; +val Scons_inject2 = thm "Scons_inject2"; +val Scons_inject = thm "Scons_inject"; +val Scons_Scons_eq = thm "Scons_Scons_eq"; +val Scons_not_Leaf = thm "Scons_not_Leaf"; +val Leaf_not_Scons = thm "Leaf_not_Scons"; +val Scons_not_Numb = thm "Scons_not_Numb"; +val Numb_not_Scons = thm "Numb_not_Scons"; +val Leaf_not_Numb = thm "Leaf_not_Numb"; +val Numb_not_Leaf = thm "Numb_not_Leaf"; +val ndepth_K0 = thm "ndepth_K0"; +val ndepth_Push_Node_aux = thm "ndepth_Push_Node_aux"; +val ndepth_Push_Node = thm "ndepth_Push_Node"; +val ntrunc_0 = thm "ntrunc_0"; +val ntrunc_Atom = thm "ntrunc_Atom"; +val ntrunc_Leaf = thm "ntrunc_Leaf"; +val ntrunc_Numb = thm "ntrunc_Numb"; +val ntrunc_Scons = thm "ntrunc_Scons"; +val ntrunc_one_In0 = thm "ntrunc_one_In0"; +val ntrunc_In0 = thm "ntrunc_In0"; +val ntrunc_one_In1 = thm "ntrunc_one_In1"; +val ntrunc_In1 = thm "ntrunc_In1"; +val uprodI = thm "uprodI"; +val uprodE = thm "uprodE"; +val uprodE2 = thm "uprodE2"; +val usum_In0I = thm "usum_In0I"; +val usum_In1I = thm "usum_In1I"; +val usumE = thm "usumE"; +val In0_not_In1 = thm "In0_not_In1"; +val In1_not_In0 = thm "In1_not_In0"; +val In0_inject = thm "In0_inject"; +val In1_inject = thm "In1_inject"; +val In0_eq = thm "In0_eq"; +val In1_eq = thm "In1_eq"; +val inj_In0 = thm "inj_In0"; +val inj_In1 = thm "inj_In1"; +val Lim_inject = thm "Lim_inject"; +val ntrunc_subsetI = thm "ntrunc_subsetI"; +val ntrunc_subsetD = thm "ntrunc_subsetD"; +val ntrunc_equality = thm "ntrunc_equality"; +val ntrunc_o_equality = thm "ntrunc_o_equality"; +val uprod_mono = thm "uprod_mono"; +val usum_mono = thm "usum_mono"; +val Scons_mono = thm "Scons_mono"; +val In0_mono = thm "In0_mono"; +val In1_mono = thm "In1_mono"; +val Split = thm "Split"; +val Case_In0 = thm "Case_In0"; +val Case_In1 = thm "Case_In1"; +val ntrunc_UN1 = thm "ntrunc_UN1"; +val Scons_UN1_x = thm "Scons_UN1_x"; +val Scons_UN1_y = thm "Scons_UN1_y"; +val In0_UN1 = thm "In0_UN1"; +val In1_UN1 = thm "In1_UN1"; +val dprodI = thm "dprodI"; +val dprodE = thm "dprodE"; +val dsum_In0I = thm "dsum_In0I"; +val dsum_In1I = thm "dsum_In1I"; +val dsumE = thm "dsumE"; +val dprod_mono = thm "dprod_mono"; +val dsum_mono = thm "dsum_mono"; +val dprod_Sigma = thm "dprod_Sigma"; +val dprod_subset_Sigma = thm "dprod_subset_Sigma"; +val dprod_subset_Sigma2 = thm "dprod_subset_Sigma2"; +val dsum_Sigma = thm "dsum_Sigma"; +val dsum_subset_Sigma = thm "dsum_subset_Sigma"; +val Domain_dprod = thm "Domain_dprod"; +val Domain_dsum = thm "Domain_dsum"; +*} + end diff -r 24aff9e3de3f -r aa785cea8fff src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Dec 08 15:15:49 2004 +0100 +++ b/src/HOL/IsaMakefile Thu Dec 09 12:03:06 2004 +0100 @@ -81,7 +81,7 @@ $(SRC)/TFL/dcterm.ML $(SRC)/TFL/post.ML $(SRC)/TFL/rules.ML \ $(SRC)/TFL/tfl.ML $(SRC)/TFL/thms.ML $(SRC)/TFL/thry.ML \ $(SRC)/TFL/usyntax.ML $(SRC)/TFL/utils.ML \ - Datatype.thy Datatype_Universe.ML Datatype_Universe.thy \ + Datatype.thy Datatype_Universe.thy \ Divides.thy Equiv_Relations.thy Extraction.thy Finite_Set.ML Finite_Set.thy \ Fun.thy Gfp.thy Hilbert_Choice.thy HOL.ML \ HOL.thy HOL_lemmas.ML Inductive.thy Infinite_Set.thy Integ/Numeral.thy \