# HG changeset patch # User paulson # Date 1102607146 -3600 # Node ID 797ed46d724b038eae090c96e8e8a83123f31539 # Parent 87f78411f7c900a05083eed704f61b7eda2f8ef4 converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global diff -r 87f78411f7c9 -r 797ed46d724b TODO --- a/TODO Thu Dec 09 15:49:40 2004 +0100 +++ b/TODO Thu Dec 09 16:45:46 2004 +0100 @@ -14,11 +14,9 @@ who is interested) - eliminate the last remaining old-style theories (Larry): - Datatype_Universe.ML HOL_lemmas.ML NatArith.ML Relation_Power.ML - Sum_Type.ML - update or remove ex/MT (Larry? Tobias?) diff -r 87f78411f7c9 -r 797ed46d724b src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Dec 09 15:49:40 2004 +0100 +++ b/src/HOL/IsaMakefile Thu Dec 09 16:45:46 2004 +0100 @@ -98,7 +98,7 @@ Record.thy Relation.ML Relation.thy Relation_Power.ML \ Relation_Power.thy LOrder.thy OrderedGroup.thy OrderedGroup.ML Ring_and_Field.thy\ Set.ML Set.thy SetInterval.ML SetInterval.thy \ - Sum_Type.ML Sum_Type.thy Tools/datatype_abs_proofs.ML Tools/datatype_aux.ML \ + Sum_Type.thy Tools/datatype_abs_proofs.ML Tools/datatype_aux.ML \ Tools/datatype_codegen.ML Tools/datatype_package.ML Tools/datatype_prop.ML \ Tools/datatype_realizer.ML Tools/datatype_rep_proofs.ML \ Tools/inductive_codegen.ML Tools/inductive_package.ML Tools/inductive_realizer.ML \ diff -r 87f78411f7c9 -r 797ed46d724b src/HOL/Sum_Type.ML --- a/src/HOL/Sum_Type.ML Thu Dec 09 15:49:40 2004 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,178 +0,0 @@ -(* Title: HOL/Sum_Type.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1991 University of Cambridge - -The disjoint sum of two types -*) - -(** Inl_Rep and Inr_Rep: Representations of the constructors **) - -(*This counts as a non-emptiness result for admitting 'a+'b as a type*) -Goalw [Sum_def] "Inl_Rep(a) : Sum"; -by (EVERY1 [rtac CollectI, rtac disjI1, rtac exI, rtac refl]); -qed "Inl_RepI"; - -Goalw [Sum_def] "Inr_Rep(b) : Sum"; -by (EVERY1 [rtac CollectI, rtac disjI2, rtac exI, rtac refl]); -qed "Inr_RepI"; - -Goal "inj_on Abs_Sum Sum"; -by (rtac inj_on_inverseI 1); -by (etac Abs_Sum_inverse 1); -qed "inj_on_Abs_Sum"; - -(** Distinctness of Inl and Inr **) - -Goalw [Inl_Rep_def, Inr_Rep_def] "Inl_Rep(a) ~= Inr_Rep(b)"; -by (EVERY1 [rtac notI, - etac (fun_cong RS fun_cong RS fun_cong RS iffE), - rtac (notE RS ccontr), etac (mp RS conjunct2), - REPEAT o (ares_tac [refl,conjI]) ]); -qed "Inl_Rep_not_Inr_Rep"; - -Goalw [Inl_def,Inr_def] "Inl(a) ~= Inr(b)"; -by (rtac (inj_on_Abs_Sum RS inj_on_contraD) 1); -by (rtac Inl_Rep_not_Inr_Rep 1); -by (rtac Inl_RepI 1); -by (rtac Inr_RepI 1); -qed "Inl_not_Inr"; - -bind_thm ("Inr_not_Inl", Inl_not_Inr RS not_sym); - -AddIffs [Inl_not_Inr, Inr_not_Inl]; - -bind_thm ("Inl_neq_Inr", Inl_not_Inr RS notE); -bind_thm ("Inr_neq_Inl", sym RS Inl_neq_Inr); - - -(** Injectiveness of Inl and Inr **) - -Goalw [Inl_Rep_def] "Inl_Rep(a) = Inl_Rep(c) ==> a=c"; -by (etac (fun_cong RS fun_cong RS fun_cong RS iffE) 1); -by (Blast_tac 1); -qed "Inl_Rep_inject"; - -Goalw [Inr_Rep_def] "Inr_Rep(b) = Inr_Rep(d) ==> b=d"; -by (etac (fun_cong RS fun_cong RS fun_cong RS iffE) 1); -by (Blast_tac 1); -qed "Inr_Rep_inject"; - -Goalw [Inl_def] "inj(Inl)"; -by (rtac injI 1); -by (etac (inj_on_Abs_Sum RS inj_onD RS Inl_Rep_inject) 1); -by (rtac Inl_RepI 1); -by (rtac Inl_RepI 1); -qed "inj_Inl"; -bind_thm ("Inl_inject", inj_Inl RS injD); - -Goalw [Inr_def] "inj(Inr)"; -by (rtac injI 1); -by (etac (inj_on_Abs_Sum RS inj_onD RS Inr_Rep_inject) 1); -by (rtac Inr_RepI 1); -by (rtac Inr_RepI 1); -qed "inj_Inr"; -bind_thm ("Inr_inject", inj_Inr RS injD); - -Goal "(Inl(x)=Inl(y)) = (x=y)"; -by (blast_tac (claset() addSDs [Inl_inject]) 1); -qed "Inl_eq"; - -Goal "(Inr(x)=Inr(y)) = (x=y)"; -by (blast_tac (claset() addSDs [Inr_inject]) 1); -qed "Inr_eq"; - -AddIffs [Inl_eq, Inr_eq]; - -(*** Rules for the disjoint sum of two SETS ***) - -(** Introduction rules for the injections **) - -Goalw [sum_def] "a : A ==> Inl(a) : A <+> B"; -by (Blast_tac 1); -qed "InlI"; - -Goalw [sum_def] "b : B ==> Inr(b) : A <+> B"; -by (Blast_tac 1); -qed "InrI"; - -(** Elimination rules **) - -val major::prems = Goalw [sum_def] - "[| u: A <+> B; \ -\ !!x. [| x:A; u=Inl(x) |] ==> P; \ -\ !!y. [| y:B; u=Inr(y) |] ==> P \ -\ |] ==> P"; -by (rtac (major RS UnE) 1); -by (REPEAT (rtac refl 1 - ORELSE eresolve_tac (prems@[imageE,ssubst]) 1)); -qed "PlusE"; - - -AddSIs [InlI, InrI]; -AddSEs [PlusE]; - - -(** Exhaustion rule for sums -- a degenerate form of induction **) - -val prems = Goalw [Inl_def,Inr_def] - "[| !!x::'a. s = Inl(x) ==> P; !!y::'b. s = Inr(y) ==> P \ -\ |] ==> P"; -by (rtac (rewrite_rule [Sum_def] Rep_Sum RS CollectE) 1); -by (REPEAT (eresolve_tac [disjE,exE] 1 - ORELSE EVERY1 [resolve_tac prems, - etac subst, - rtac (Rep_Sum_inverse RS sym)])); -qed "sumE"; - -val prems = Goal "[| !!x. P (Inl x); !!x. P (Inr x) |] ==> P x"; -by (res_inst_tac [("s","x")] sumE 1); -by (ALLGOALS (hyp_subst_tac THEN' (resolve_tac prems))); -qed "sum_induct"; - - -(** Rules for the Part primitive **) - -Goalw [Part_def] "[| a : A; a=h(b) |] ==> a : Part A h"; -by (Blast_tac 1); -qed "Part_eqI"; - -bind_thm ("PartI", refl RSN (2,Part_eqI)); - -val major::prems = Goalw [Part_def] - "[| a : Part A h; !!z. [| a : A; a=h(z) |] ==> P \ -\ |] ==> P"; -by (rtac (major RS IntE) 1); -by (etac CollectE 1); -by (etac exE 1); -by (REPEAT (ares_tac prems 1)); -qed "PartE"; - -AddIs [Part_eqI]; -AddSEs [PartE]; - -Goalw [Part_def] "Part A h <= A"; -by (rtac Int_lower1 1); -qed "Part_subset"; - -Goal "A<=B ==> Part A h <= Part B h"; -by (Blast_tac 1); -qed "Part_mono"; - -bind_thms ("basic_monos", basic_monos @ [Part_mono]); - -Goalw [Part_def] "a : Part A h ==> a : A"; -by (etac IntD1 1); -qed "PartD1"; - -Goal "Part A (%x. x) = A"; -by (Blast_tac 1); -qed "Part_id"; - -Goal "Part (A Int B) h = (Part A h) Int (Part B h)"; -by (Blast_tac 1); -qed "Part_Int"; - -Goal "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}"; -by (Blast_tac 1); -qed "Part_Collect"; diff -r 87f78411f7c9 -r 797ed46d724b src/HOL/Sum_Type.thy --- a/src/HOL/Sum_Type.thy Thu Dec 09 15:49:40 2004 +0100 +++ b/src/HOL/Sum_Type.thy Thu Dec 09 16:45:46 2004 +0100 @@ -2,47 +2,221 @@ ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge - -The disjoint sum of two types. *) -Sum_Type = Product_Type + +header{*The Disjoint Sum of Two Types*} -(* type definition *) +theory Sum_Type +imports Product_Type +begin + +text{*The representations of the two injections*} constdefs - Inl_Rep :: ['a, 'a, 'b, bool] => bool + Inl_Rep :: "['a, 'a, 'b, bool] => bool" "Inl_Rep == (%a. %x y p. x=a & p)" - Inr_Rep :: ['b, 'a, 'b, bool] => bool + Inr_Rep :: "['b, 'a, 'b, bool] => bool" "Inr_Rep == (%b. %x y p. y=b & ~p)" + global typedef (Sum) ('a, 'b) "+" (infixr 10) = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}" - - -(* abstract constants and syntax *) - -consts - Inl :: "'a => 'a + 'b" - Inr :: "'b => 'a + 'b" - - (*disjoint sum for sets; the operator + is overloaded with wrong type!*) - Plus :: "['a set, 'b set] => ('a + 'b) set" (infixr "<+>" 65) - Part :: ['a set, 'b => 'a] => 'a set + by auto local -defs - Inl_def "Inl == (%a. Abs_Sum(Inl_Rep(a)))" - Inr_def "Inr == (%b. Abs_Sum(Inr_Rep(b)))" + +text{*abstract constants and syntax*} + +constdefs + Inl :: "'a => 'a + 'b" + "Inl == (%a. Abs_Sum(Inl_Rep(a)))" + + Inr :: "'b => 'a + 'b" + "Inr == (%b. Abs_Sum(Inr_Rep(b)))" + + Plus :: "['a set, 'b set] => ('a + 'b) set" (infixr "<+>" 65) + "A <+> B == (Inl`A) Un (Inr`B)" + --{*disjoint sum for sets; the operator + is overloaded with wrong type!*} + + Part :: "['a set, 'b => 'a] => 'a set" + "Part A h == A Int {x. ? z. x = h(z)}" + --{*for selecting out the components of a mutually recursive definition*} + + + +(** Inl_Rep and Inr_Rep: Representations of the constructors **) + +(*This counts as a non-emptiness result for admitting 'a+'b as a type*) +lemma Inl_RepI: "Inl_Rep(a) : Sum" +by (auto simp add: Sum_def) + +lemma Inr_RepI: "Inr_Rep(b) : Sum" +by (auto simp add: Sum_def) + +lemma inj_on_Abs_Sum: "inj_on Abs_Sum Sum" +apply (rule inj_on_inverseI) +apply (erule Abs_Sum_inverse) +done + +subsection{*Freeness Properties for @{term Inl} and @{term Inr}*} + +text{*Distinctness*} + +lemma Inl_Rep_not_Inr_Rep: "Inl_Rep(a) ~= Inr_Rep(b)" +by (auto simp add: Inl_Rep_def Inr_Rep_def expand_fun_eq) + +lemma Inl_not_Inr [iff]: "Inl(a) ~= Inr(b)" +apply (simp add: Inl_def Inr_def) +apply (rule inj_on_Abs_Sum [THEN inj_on_contraD]) +apply (rule Inl_Rep_not_Inr_Rep) +apply (rule Inl_RepI) +apply (rule Inr_RepI) +done + +lemmas Inr_not_Inl = Inl_not_Inr [THEN not_sym, standard, iff] + +lemmas Inl_neq_Inr = Inl_not_Inr [THEN notE, standard] +lemmas Inr_neq_Inl = sym [THEN Inl_neq_Inr, standard] + + +text{*Injectiveness*} + +lemma Inl_Rep_inject: "Inl_Rep(a) = Inl_Rep(c) ==> a=c" +by (auto simp add: Inl_Rep_def expand_fun_eq) + +lemma Inr_Rep_inject: "Inr_Rep(b) = Inr_Rep(d) ==> b=d" +by (auto simp add: Inr_Rep_def expand_fun_eq) + +lemma inj_Inl: "inj(Inl)" +apply (simp add: Inl_def) +apply (rule inj_onI) +apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inl_Rep_inject]) +apply (rule Inl_RepI) +apply (rule Inl_RepI) +done +lemmas Inl_inject = inj_Inl [THEN injD, standard] + +lemma inj_Inr: "inj(Inr)" +apply (simp add: Inr_def) +apply (rule inj_onI) +apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inr_Rep_inject]) +apply (rule Inr_RepI) +apply (rule Inr_RepI) +done + +lemmas Inr_inject = inj_Inr [THEN injD, standard] + +lemma Inl_eq [iff]: "(Inl(x)=Inl(y)) = (x=y)" +by (blast dest!: Inl_inject) + +lemma Inr_eq [iff]: "(Inr(x)=Inr(y)) = (x=y)" +by (blast dest!: Inr_inject) + + +subsection{*The Disjoint Sum of Sets*} + +(** Introduction rules for the injections **) - sum_def "A <+> B == (Inl`A) Un (Inr`B)" +lemma InlI [intro!]: "a : A ==> Inl(a) : A <+> B" +by (simp add: Plus_def) + +lemma InrI [intro!]: "b : B ==> Inr(b) : A <+> B" +by (simp add: Plus_def) + +(** Elimination rules **) + +lemma PlusE [elim!]: + "[| u: A <+> B; + !!x. [| x:A; u=Inl(x) |] ==> P; + !!y. [| y:B; u=Inr(y) |] ==> P + |] ==> P" +by (auto simp add: Plus_def) + + + +text{*Exhaustion rule for sums, a degenerate form of induction*} +lemma sumE: + "[| !!x::'a. s = Inl(x) ==> P; !!y::'b. s = Inr(y) ==> P + |] ==> P" +apply (rule Abs_Sum_cases [of s]) +apply (auto simp add: Sum_def Inl_def Inr_def) +done + +lemma sum_induct: "[| !!x. P (Inl x); !!x. P (Inr x) |] ==> P x" +by (rule sumE [of x], auto) + + +subsection{*The @{term Part} Primitive*} + +lemma Part_eqI [intro]: "[| a : A; a=h(b) |] ==> a : Part A h" +by (auto simp add: Part_def) + +lemmas PartI = Part_eqI [OF _ refl, standard] + +lemma PartE [elim!]: "[| a : Part A h; !!z. [| a : A; a=h(z) |] ==> P |] ==> P" +by (auto simp add: Part_def) + + +lemma Part_subset: "Part A h <= A" +by (auto simp add: Part_def) + +lemma Part_mono: "A<=B ==> Part A h <= Part B h" +by blast + +lemmas basic_monos = basic_monos Part_mono - (*for selecting out the components of a mutually recursive definition*) - Part_def "Part A h == A Int {x. ? z. x = h(z)}" +lemma PartD1: "a : Part A h ==> a : A" +by (simp add: Part_def) + +lemma Part_id: "Part A (%x. x) = A" +by blast + +lemma Part_Int: "Part (A Int B) h = (Part A h) Int (Part B h)" +by blast + +lemma Part_Collect: "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}" +by blast + +ML +{* +val Inl_RepI = thm "Inl_RepI"; +val Inr_RepI = thm "Inr_RepI"; +val inj_on_Abs_Sum = thm "inj_on_Abs_Sum"; +val Inl_Rep_not_Inr_Rep = thm "Inl_Rep_not_Inr_Rep"; +val Inl_not_Inr = thm "Inl_not_Inr"; +val Inr_not_Inl = thm "Inr_not_Inl"; +val Inl_neq_Inr = thm "Inl_neq_Inr"; +val Inr_neq_Inl = thm "Inr_neq_Inl"; +val Inl_Rep_inject = thm "Inl_Rep_inject"; +val Inr_Rep_inject = thm "Inr_Rep_inject"; +val inj_Inl = thm "inj_Inl"; +val Inl_inject = thm "Inl_inject"; +val inj_Inr = thm "inj_Inr"; +val Inr_inject = thm "Inr_inject"; +val Inl_eq = thm "Inl_eq"; +val Inr_eq = thm "Inr_eq"; +val InlI = thm "InlI"; +val InrI = thm "InrI"; +val PlusE = thm "PlusE"; +val sumE = thm "sumE"; +val sum_induct = thm "sum_induct"; +val Part_eqI = thm "Part_eqI"; +val PartI = thm "PartI"; +val PartE = thm "PartE"; +val Part_subset = thm "Part_subset"; +val Part_mono = thm "Part_mono"; +val PartD1 = thm "PartD1"; +val Part_id = thm "Part_id"; +val Part_Int = thm "Part_Int"; +val Part_Collect = thm "Part_Collect"; + +val basic_monos = thms "basic_monos"; +*} + end diff -r 87f78411f7c9 -r 797ed46d724b src/HOL/Tools/datatype_rep_proofs.ML --- a/src/HOL/Tools/datatype_rep_proofs.ML Thu Dec 09 15:49:40 2004 +0100 +++ b/src/HOL/Tools/datatype_rep_proofs.ML Thu Dec 09 16:45:46 2004 +0100 @@ -103,9 +103,9 @@ val Type (_, [T1, T2]) = T in if i <= n2 then - Const ("Inl", T1 --> T) $ (mk_inj' T1 n2 i) + Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i) else - Const ("Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) + Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) end in mk_inj' sumT (length leafTs) (1 + find_index_eq T' leafTs) end; diff -r 87f78411f7c9 -r 797ed46d724b src/HOL/Tools/inductive_package.ML --- a/src/HOL/Tools/inductive_package.ML Thu Dec 09 15:49:40 2004 +0100 +++ b/src/HOL/Tools/inductive_package.ML Thu Dec 09 16:45:46 2004 +0100 @@ -212,9 +212,9 @@ val Type (_, [T1, T2]) = T in if i <= n2 then - Const ("Inl", T1 --> T) $ (mk_inj' T1 n2 i) + Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i) else - Const ("Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) + Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) end in mk_inj' sumT (length cs) (1 + find_index_eq c cs) end;