nipkow@10213: (* Title: HOL/Product_Type.thy nipkow@10213: Author: Lawrence C Paulson, Cambridge University Computer Laboratory nipkow@10213: Copyright 1992 University of Cambridge wenzelm@11777: *) nipkow@10213: wenzelm@11838: header {* Cartesian products *} nipkow@10213: nipkow@15131: theory Product_Type haftmann@33959: imports Typedef Inductive Fun wenzelm@46950: keywords "inductive_set" "coinductive_set" :: thy_decl nipkow@15131: begin wenzelm@11838: haftmann@24699: subsection {* @{typ bool} is a datatype *} haftmann@24699: haftmann@27104: rep_datatype True False by (auto intro: bool_induct) haftmann@24699: haftmann@24699: declare case_split [cases type: bool] haftmann@24699: -- "prefer plain propositional version" haftmann@24699: haftmann@28346: lemma haftmann@38857: shows [code]: "HOL.equal False P \ \ P" haftmann@38857: and [code]: "HOL.equal True P \ P" haftmann@46630: and [code]: "HOL.equal P False \ \ P" haftmann@38857: and [code]: "HOL.equal P True \ P" haftmann@38857: and [code nbe]: "HOL.equal P P \ True" haftmann@38857: by (simp_all add: equal) haftmann@25534: haftmann@43654: lemma If_case_cert: haftmann@43654: assumes "CASE \ (\b. If b f g)" haftmann@43654: shows "(CASE True \ f) &&& (CASE False \ g)" haftmann@43654: using assms by simp_all haftmann@43654: haftmann@43654: setup {* haftmann@43654: Code.add_case @{thm If_case_cert} haftmann@43654: *} haftmann@43654: haftmann@38857: code_const "HOL.equal \ bool \ bool \ bool" haftmann@39272: (Haskell infix 4 "==") haftmann@25534: haftmann@38857: code_instance bool :: equal haftmann@25534: (Haskell -) haftmann@24699: haftmann@26358: haftmann@37166: subsection {* The @{text unit} type *} wenzelm@11838: wenzelm@49834: typedef unit = "{True}" wenzelm@45694: by auto wenzelm@11838: wenzelm@45694: definition Unity :: unit ("'(')") wenzelm@45694: where "() = Abs_unit True" wenzelm@11838: blanchet@35828: lemma unit_eq [no_atp]: "u = ()" huffman@40590: by (induct u) (simp add: Unity_def) wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: Simplification procedure for @{thm [source] unit_eq}. Cannot use wenzelm@11838: this rule directly --- it loops! wenzelm@11838: *} wenzelm@11838: wenzelm@43594: simproc_setup unit_eq ("x::unit") = {* wenzelm@43594: fn _ => fn _ => fn ct => wenzelm@43594: if HOLogic.is_unit (term_of ct) then NONE wenzelm@43594: else SOME (mk_meta_eq @{thm unit_eq}) wenzelm@11838: *} wenzelm@11838: haftmann@27104: rep_datatype "()" by simp haftmann@24699: wenzelm@11838: lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()" wenzelm@11838: by simp wenzelm@11838: wenzelm@11838: lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P" wenzelm@11838: by (rule triv_forall_equality) wenzelm@11838: wenzelm@11838: text {* wenzelm@43594: This rewrite counters the effect of simproc @{text unit_eq} on @{term wenzelm@11838: [source] "%u::unit. f u"}, replacing it by @{term [source] wenzelm@11838: f} rather than by @{term [source] "%u. f ()"}. wenzelm@11838: *} wenzelm@11838: haftmann@43866: lemma unit_abs_eta_conv [simp, no_atp]: "(%u::unit. f ()) = f" wenzelm@11838: by (rule ext) simp nipkow@10213: haftmann@43866: lemma UNIV_unit [no_atp]: haftmann@43866: "UNIV = {()}" by auto haftmann@43866: haftmann@30924: instantiation unit :: default haftmann@30924: begin haftmann@30924: haftmann@30924: definition "default = ()" haftmann@30924: haftmann@30924: instance .. haftmann@30924: haftmann@30924: end nipkow@10213: haftmann@28562: lemma [code]: haftmann@38857: "HOL.equal (u\unit) v \ True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+ haftmann@26358: haftmann@26358: code_type unit haftmann@26358: (SML "unit") haftmann@26358: (OCaml "unit") haftmann@26358: (Haskell "()") haftmann@34886: (Scala "Unit") haftmann@26358: haftmann@37166: code_const Unity haftmann@37166: (SML "()") haftmann@37166: (OCaml "()") haftmann@37166: (Haskell "()") haftmann@37166: (Scala "()") haftmann@37166: haftmann@38857: code_instance unit :: equal haftmann@26358: (Haskell -) haftmann@26358: haftmann@38857: code_const "HOL.equal \ unit \ unit \ bool" haftmann@39272: (Haskell infix 4 "==") haftmann@26358: haftmann@26358: code_reserved SML haftmann@26358: unit haftmann@26358: haftmann@26358: code_reserved OCaml haftmann@26358: unit haftmann@26358: haftmann@34886: code_reserved Scala haftmann@34886: Unit haftmann@34886: haftmann@26358: haftmann@37166: subsection {* The product type *} nipkow@10213: haftmann@37166: subsubsection {* Type definition *} haftmann@37166: haftmann@37166: definition Pair_Rep :: "'a \ 'b \ 'a \ 'b \ bool" where haftmann@26358: "Pair_Rep a b = (\x y. x = a \ y = b)" nipkow@10213: wenzelm@45696: definition "prod = {f. \a b. f = Pair_Rep (a\'a) (b\'b)}" wenzelm@45696: wenzelm@49834: typedef ('a, 'b) prod (infixr "*" 20) = "prod :: ('a \ 'b \ bool) set" wenzelm@45696: unfolding prod_def by auto nipkow@10213: wenzelm@35427: type_notation (xsymbols) haftmann@37678: "prod" ("(_ \/ _)" [21, 20] 20) wenzelm@35427: type_notation (HTML output) haftmann@37678: "prod" ("(_ \/ _)" [21, 20] 20) nipkow@10213: haftmann@37389: definition Pair :: "'a \ 'b \ 'a \ 'b" where haftmann@37389: "Pair a b = Abs_prod (Pair_Rep a b)" haftmann@37166: haftmann@37678: rep_datatype Pair proof - haftmann@37166: fix P :: "'a \ 'b \ bool" and p haftmann@37166: assume "\a b. P (Pair a b)" haftmann@37389: then show "P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) haftmann@37166: next haftmann@37166: fix a c :: 'a and b d :: 'b haftmann@37166: have "Pair_Rep a b = Pair_Rep c d \ a = c \ b = d" nipkow@39302: by (auto simp add: Pair_Rep_def fun_eq_iff) haftmann@37389: moreover have "Pair_Rep a b \ prod" and "Pair_Rep c d \ prod" haftmann@37389: by (auto simp add: prod_def) haftmann@37166: ultimately show "Pair a b = Pair c d \ a = c \ b = d" haftmann@37389: by (simp add: Pair_def Abs_prod_inject) haftmann@37166: qed haftmann@37166: blanchet@37704: declare prod.simps(2) [nitpick_simp del] blanchet@37704: huffman@40929: declare prod.weak_case_cong [cong del] haftmann@37411: haftmann@37166: haftmann@37166: subsubsection {* Tuple syntax *} haftmann@37166: haftmann@37591: abbreviation (input) split :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where haftmann@37591: "split \ prod_case" wenzelm@19535: wenzelm@11777: text {* wenzelm@11777: Patterns -- extends pre-defined type @{typ pttrn} used in wenzelm@11777: abstractions. wenzelm@11777: *} nipkow@10213: wenzelm@41229: nonterminal tuple_args and patterns nipkow@10213: nipkow@10213: syntax nipkow@10213: "_tuple" :: "'a => tuple_args => 'a * 'b" ("(1'(_,/ _'))") nipkow@10213: "_tuple_arg" :: "'a => tuple_args" ("_") nipkow@10213: "_tuple_args" :: "'a => tuple_args => tuple_args" ("_,/ _") oheimb@11025: "_pattern" :: "[pttrn, patterns] => pttrn" ("'(_,/ _')") oheimb@11025: "" :: "pttrn => patterns" ("_") oheimb@11025: "_patterns" :: "[pttrn, patterns] => patterns" ("_,/ _") nipkow@10213: nipkow@10213: translations wenzelm@35115: "(x, y)" == "CONST Pair x y" nipkow@10213: "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))" haftmann@37591: "%(x, y, zs). b" == "CONST prod_case (%x (y, zs). b)" haftmann@37591: "%(x, y). b" == "CONST prod_case (%x y. b)" wenzelm@35115: "_abs (CONST Pair x y) t" => "%(x, y). t" haftmann@37166: -- {* The last rule accommodates tuples in `case C ... (x,y) ... => ...' haftmann@37166: The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *} nipkow@10213: wenzelm@35115: (*reconstruct pattern from (nested) splits, avoiding eta-contraction of body; wenzelm@35115: works best with enclosing "let", if "let" does not avoid eta-contraction*) schirmer@14359: print_translation {* wenzelm@35115: let wenzelm@35115: fun split_tr' [Abs (x, T, t as (Abs abs))] = wenzelm@35115: (* split (%x y. t) => %(x,y) t *) wenzelm@35115: let wenzelm@42284: val (y, t') = Syntax_Trans.atomic_abs_tr' abs; wenzelm@42284: val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); wenzelm@35115: in wenzelm@35115: Syntax.const @{syntax_const "_abs"} $ wenzelm@35115: (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t'' wenzelm@35115: end haftmann@37591: | split_tr' [Abs (x, T, (s as Const (@{const_syntax prod_case}, _) $ t))] = wenzelm@35115: (* split (%x. (split (%y z. t))) => %(x,y,z). t *) wenzelm@35115: let wenzelm@35115: val Const (@{syntax_const "_abs"}, _) $ wenzelm@35115: (Const (@{syntax_const "_pattern"}, _) $ y $ z) $ t' = split_tr' [t]; wenzelm@42284: val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); wenzelm@35115: in wenzelm@35115: Syntax.const @{syntax_const "_abs"} $ wenzelm@35115: (Syntax.const @{syntax_const "_pattern"} $ x' $ wenzelm@35115: (Syntax.const @{syntax_const "_patterns"} $ y $ z)) $ t'' wenzelm@35115: end haftmann@37591: | split_tr' [Const (@{const_syntax prod_case}, _) $ t] = wenzelm@35115: (* split (split (%x y z. t)) => %((x, y), z). t *) wenzelm@35115: split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *) wenzelm@35115: | split_tr' [Const (@{syntax_const "_abs"}, _) $ x_y $ Abs abs] = wenzelm@35115: (* split (%pttrn z. t) => %(pttrn,z). t *) wenzelm@42284: let val (z, t) = Syntax_Trans.atomic_abs_tr' abs in wenzelm@35115: Syntax.const @{syntax_const "_abs"} $ wenzelm@35115: (Syntax.const @{syntax_const "_pattern"} $ x_y $ z) $ t wenzelm@35115: end wenzelm@35115: | split_tr' _ = raise Match; haftmann@37591: in [(@{const_syntax prod_case}, split_tr')] end schirmer@14359: *} schirmer@14359: schirmer@15422: (* print "split f" as "\(x,y). f x y" and "split (\x. f x)" as "\(x,y). f x y" *) schirmer@15422: typed_print_translation {* schirmer@15422: let wenzelm@42247: fun split_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match wenzelm@42247: | split_guess_names_tr' T [Abs (x, xT, t)] = schirmer@15422: (case (head_of t) of haftmann@37591: Const (@{const_syntax prod_case}, _) => raise Match wenzelm@35115: | _ => wenzelm@35115: let wenzelm@35115: val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; wenzelm@42284: val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0); wenzelm@42284: val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, xT, t'); wenzelm@35115: in wenzelm@35115: Syntax.const @{syntax_const "_abs"} $ wenzelm@35115: (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t'' wenzelm@35115: end) wenzelm@42247: | split_guess_names_tr' T [t] = wenzelm@35115: (case head_of t of haftmann@37591: Const (@{const_syntax prod_case}, _) => raise Match wenzelm@35115: | _ => wenzelm@35115: let wenzelm@35115: val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; wenzelm@42284: val (y, t') = wenzelm@42284: Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0); wenzelm@42284: val (x', t'') = Syntax_Trans.atomic_abs_tr' ("x", xT, t'); wenzelm@35115: in wenzelm@35115: Syntax.const @{syntax_const "_abs"} $ wenzelm@35115: (Syntax.const @{syntax_const "_pattern"} $ x' $ y) $ t'' wenzelm@35115: end) wenzelm@42247: | split_guess_names_tr' _ _ = raise Match; haftmann@37591: in [(@{const_syntax prod_case}, split_guess_names_tr')] end schirmer@15422: *} schirmer@15422: nipkow@42059: (* Force eta-contraction for terms of the form "Q A (%p. prod_case P p)" nipkow@42059: where Q is some bounded quantifier or set operator. nipkow@42059: Reason: the above prints as "Q p : A. case p of (x,y) \ P x y" nipkow@42059: whereas we want "Q (x,y):A. P x y". nipkow@42059: Otherwise prevent eta-contraction. nipkow@42059: *) nipkow@42059: print_translation {* nipkow@42059: let nipkow@42059: fun contract Q f ts = nipkow@42059: case ts of nipkow@42059: [A, Abs(_, _, (s as Const (@{const_syntax prod_case},_) $ t) $ Bound 0)] wenzelm@42083: => if Term.is_dependent t then f ts else Syntax.const Q $ A $ s nipkow@42059: | _ => f ts; nipkow@42059: fun contract2 (Q,f) = (Q, contract Q f); nipkow@42059: val pairs = wenzelm@42284: [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Ball} @{syntax_const "_Ball"}, wenzelm@42284: Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax Bex} @{syntax_const "_Bex"}, wenzelm@42284: Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}, wenzelm@42284: Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}] nipkow@42059: in map contract2 pairs end nipkow@42059: *} nipkow@10213: haftmann@37166: subsubsection {* Code generator setup *} haftmann@37166: haftmann@37678: code_type prod haftmann@37166: (SML infix 2 "*") haftmann@37166: (OCaml infix 2 "*") haftmann@37166: (Haskell "!((_),/ (_))") haftmann@37166: (Scala "((_),/ (_))") haftmann@37166: haftmann@37166: code_const Pair haftmann@37166: (SML "!((_),/ (_))") haftmann@37166: (OCaml "!((_),/ (_))") haftmann@37166: (Haskell "!((_),/ (_))") haftmann@37166: (Scala "!((_),/ (_))") haftmann@37166: haftmann@38857: code_instance prod :: equal haftmann@37166: (Haskell -) haftmann@37166: haftmann@38857: code_const "HOL.equal \ 'a \ 'b \ 'a \ 'b \ bool" haftmann@39272: (Haskell infix 4 "==") haftmann@37166: haftmann@37166: haftmann@37166: subsubsection {* Fundamental operations and properties *} wenzelm@11838: haftmann@26358: lemma surj_pair [simp]: "EX x y. p = (x, y)" haftmann@37166: by (cases p) simp nipkow@10213: haftmann@37389: definition fst :: "'a \ 'b \ 'a" where haftmann@37389: "fst p = (case p of (a, b) \ a)" wenzelm@11838: haftmann@37389: definition snd :: "'a \ 'b \ 'b" where haftmann@37389: "snd p = (case p of (a, b) \ b)" wenzelm@11838: haftmann@22886: lemma fst_conv [simp, code]: "fst (a, b) = a" haftmann@37166: unfolding fst_def by simp wenzelm@11838: haftmann@22886: lemma snd_conv [simp, code]: "snd (a, b) = b" haftmann@37166: unfolding snd_def by simp oheimb@11025: haftmann@37166: code_const fst and snd haftmann@37166: (Haskell "fst" and "snd") haftmann@26358: blanchet@41792: lemma prod_case_unfold [nitpick_unfold]: "prod_case = (%c p. c (fst p) (snd p))" nipkow@39302: by (simp add: fun_eq_iff split: prod.split) haftmann@26358: wenzelm@11838: lemma fst_eqD: "fst (x, y) = a ==> x = a" wenzelm@11838: by simp wenzelm@11838: wenzelm@11838: lemma snd_eqD: "snd (x, y) = a ==> y = a" wenzelm@11838: by simp wenzelm@11838: haftmann@26358: lemma pair_collapse [simp]: "(fst p, snd p) = p" wenzelm@11838: by (cases p) simp wenzelm@11838: haftmann@26358: lemmas surjective_pairing = pair_collapse [symmetric] wenzelm@11838: huffman@44066: lemma prod_eq_iff: "s = t \ fst s = fst t \ snd s = snd t" haftmann@37166: by (cases s, cases t) simp haftmann@37166: haftmann@37166: lemma prod_eqI [intro?]: "fst p = fst q \ snd p = snd q \ p = q" huffman@44066: by (simp add: prod_eq_iff) haftmann@37166: haftmann@37166: lemma split_conv [simp, code]: "split f (a, b) = f a b" haftmann@37591: by (fact prod.cases) haftmann@37166: haftmann@37166: lemma splitI: "f a b \ split f (a, b)" haftmann@37166: by (rule split_conv [THEN iffD2]) haftmann@37166: haftmann@37166: lemma splitD: "split f (a, b) \ f a b" haftmann@37166: by (rule split_conv [THEN iffD1]) haftmann@37166: haftmann@37166: lemma split_Pair [simp]: "(\(x, y). (x, y)) = id" nipkow@39302: by (simp add: fun_eq_iff split: prod.split) haftmann@37166: haftmann@37166: lemma split_eta: "(\(x, y). f (x, y)) = f" haftmann@37166: -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *} nipkow@39302: by (simp add: fun_eq_iff split: prod.split) haftmann@37166: haftmann@37166: lemma split_comp: "split (f \ g) x = f (g (fst x)) (snd x)" haftmann@37166: by (cases x) simp haftmann@37166: haftmann@37166: lemma split_twice: "split f (split g p) = split (\x y. split f (g x y)) p" haftmann@37166: by (cases p) simp haftmann@37166: haftmann@37166: lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))" haftmann@37591: by (simp add: prod_case_unfold) haftmann@37166: haftmann@37166: lemma split_weak_cong: "p = q \ split c p = split c q" haftmann@37166: -- {* Prevents simplification of @{term c}: much faster *} huffman@40929: by (fact prod.weak_case_cong) haftmann@37166: haftmann@37166: lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g" haftmann@37166: by (simp add: split_eta) haftmann@37166: blanchet@47740: lemma split_paired_all [no_atp]: "(!!x. PROP P x) == (!!a b. PROP P (a, b))" wenzelm@11820: proof wenzelm@11820: fix a b wenzelm@11820: assume "!!x. PROP P x" wenzelm@19535: then show "PROP P (a, b)" . wenzelm@11820: next wenzelm@11820: fix x wenzelm@11820: assume "!!a b. PROP P (a, b)" wenzelm@19535: from `PROP P (fst x, snd x)` show "PROP P x" by simp wenzelm@11820: qed wenzelm@11820: wenzelm@11838: text {* wenzelm@11838: The rule @{thm [source] split_paired_all} does not work with the wenzelm@11838: Simplifier because it also affects premises in congrence rules, wenzelm@11838: where this can lead to premises of the form @{text "!!a b. ... = wenzelm@11838: ?P(a, b)"} which cannot be solved by reflexivity. wenzelm@11838: *} wenzelm@11838: haftmann@26358: lemmas split_tupled_all = split_paired_all unit_all_eq2 haftmann@26358: wenzelm@26480: ML {* wenzelm@11838: (* replace parameters of product type by individual component parameters *) wenzelm@11838: val safe_full_simp_tac = generic_simp_tac true (true, false, false); wenzelm@11838: local (* filtering with exists_paired_all is an essential optimization *) wenzelm@16121: fun exists_paired_all (Const ("all", _) $ Abs (_, T, t)) = wenzelm@11838: can HOLogic.dest_prodT T orelse exists_paired_all t wenzelm@11838: | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u wenzelm@11838: | exists_paired_all (Abs (_, _, t)) = exists_paired_all t wenzelm@11838: | exists_paired_all _ = false; wenzelm@11838: val ss = HOL_basic_ss wenzelm@26340: addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}] wenzelm@43594: addsimprocs [@{simproc unit_eq}]; wenzelm@11838: in wenzelm@11838: val split_all_tac = SUBGOAL (fn (t, i) => wenzelm@11838: if exists_paired_all t then safe_full_simp_tac ss i else no_tac); wenzelm@11838: val unsafe_split_all_tac = SUBGOAL (fn (t, i) => wenzelm@11838: if exists_paired_all t then full_simp_tac ss i else no_tac); wenzelm@11838: fun split_all th = wenzelm@26340: if exists_paired_all (Thm.prop_of th) then full_simplify ss th else th; wenzelm@11838: end; wenzelm@26340: *} wenzelm@11838: wenzelm@26340: declaration {* fn _ => wenzelm@26340: Classical.map_cs (fn cs => cs addSbefore ("split_all_tac", split_all_tac)) wenzelm@16121: *} wenzelm@11838: blanchet@47740: lemma split_paired_All [simp, no_atp]: "(ALL x. P x) = (ALL a b. P (a, b))" wenzelm@11838: -- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *} wenzelm@11838: by fast wenzelm@11838: blanchet@47740: lemma split_paired_Ex [simp, no_atp]: "(EX x. P x) = (EX a b. P (a, b))" haftmann@26358: by fast haftmann@26358: blanchet@47740: lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))" wenzelm@11838: -- {* Can't be added to simpset: loops! *} haftmann@26358: by (simp add: split_eta) wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: Simplification procedure for @{thm [source] cond_split_eta}. Using wenzelm@11838: @{thm [source] split_eta} as a rewrite rule is not general enough, wenzelm@11838: and using @{thm [source] cond_split_eta} directly would render some wenzelm@11838: existing proofs very inefficient; similarly for @{text haftmann@26358: split_beta}. haftmann@26358: *} wenzelm@11838: wenzelm@26480: ML {* wenzelm@11838: local wenzelm@35364: val cond_split_eta_ss = HOL_basic_ss addsimps @{thms cond_split_eta}; wenzelm@35364: fun Pair_pat k 0 (Bound m) = (m = k) wenzelm@35364: | Pair_pat k i (Const (@{const_name Pair}, _) $ Bound m $ t) = wenzelm@35364: i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t wenzelm@35364: | Pair_pat _ _ _ = false; wenzelm@35364: fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t wenzelm@35364: | no_args k i (t $ u) = no_args k i t andalso no_args k i u wenzelm@35364: | no_args k i (Bound m) = m < k orelse m > k + i wenzelm@35364: | no_args _ _ _ = true; wenzelm@35364: fun split_pat tp i (Abs (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE haftmann@37591: | split_pat tp i (Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t wenzelm@35364: | split_pat tp i _ = NONE; wenzelm@20044: fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] [] wenzelm@35364: (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))) wenzelm@18328: (K (simp_tac (Simplifier.inherit_context ss cond_split_eta_ss) 1))); wenzelm@11838: wenzelm@35364: fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t wenzelm@35364: | beta_term_pat k i (t $ u) = wenzelm@35364: Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u) wenzelm@35364: | beta_term_pat k i t = no_args k i t; wenzelm@35364: fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg wenzelm@35364: | eta_term_pat _ _ _ = false; wenzelm@11838: fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t) wenzelm@35364: | subst arg k i (t $ u) = wenzelm@35364: if Pair_pat k i (t $ u) then incr_boundvars k arg wenzelm@35364: else (subst arg k i t $ subst arg k i u) wenzelm@35364: | subst arg k i t = t; wenzelm@43595: in haftmann@37591: fun beta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t) $ arg) = wenzelm@11838: (case split_pat beta_term_pat 1 t of wenzelm@35364: SOME (i, f) => SOME (metaeq ss s (subst arg 0 i f)) skalberg@15531: | NONE => NONE) wenzelm@35364: | beta_proc _ _ = NONE; haftmann@37591: fun eta_proc ss (s as Const (@{const_name prod_case}, _) $ Abs (_, _, t)) = wenzelm@11838: (case split_pat eta_term_pat 1 t of wenzelm@35364: SOME (_, ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end)) skalberg@15531: | NONE => NONE) wenzelm@35364: | eta_proc _ _ = NONE; wenzelm@11838: end; wenzelm@11838: *} wenzelm@43595: simproc_setup split_beta ("split f z") = {* fn _ => fn ss => fn ct => beta_proc ss (term_of ct) *} wenzelm@43595: simproc_setup split_eta ("split f") = {* fn _ => fn ss => fn ct => eta_proc ss (term_of ct) *} wenzelm@11838: berghofe@26798: lemma split_beta [mono]: "(%(x, y). P x y) z = P (fst z) (snd z)" wenzelm@11838: by (subst surjective_pairing, rule split_conv) wenzelm@11838: blanchet@35828: lemma split_split [no_atp]: "R(split c p) = (ALL x y. p = (x, y) --> R(c x y))" wenzelm@11838: -- {* For use with @{text split} and the Simplifier. *} paulson@15481: by (insert surj_pair [of p], clarify, simp) wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: @{thm [source] split_split} could be declared as @{text "[split]"} wenzelm@11838: done after the Splitter has been speeded up significantly; wenzelm@11838: precompute the constants involved and don't do anything unless the wenzelm@11838: current goal contains one of those constants. wenzelm@11838: *} wenzelm@11838: blanchet@35828: lemma split_split_asm [no_atp]: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))" paulson@14208: by (subst split_split, simp) wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: \medskip @{term split} used as a logical connective or set former. wenzelm@11838: wenzelm@11838: \medskip These rules are for use with @{text blast}; could instead huffman@40929: call @{text simp} using @{thm [source] prod.split} as rewrite. *} wenzelm@11838: wenzelm@11838: lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p" wenzelm@11838: apply (simp only: split_tupled_all) wenzelm@11838: apply (simp (no_asm_simp)) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> split c p x" wenzelm@11838: apply (simp only: split_tupled_all) wenzelm@11838: apply (simp (no_asm_simp)) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma splitE: "split c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" haftmann@37591: by (induct p) auto wenzelm@11838: wenzelm@11838: lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" haftmann@37591: by (induct p) auto wenzelm@11838: wenzelm@11838: lemma splitE2: wenzelm@11838: "[| Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R" wenzelm@11838: proof - wenzelm@11838: assume q: "Q (split P z)" wenzelm@11838: assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R" wenzelm@11838: show R wenzelm@11838: apply (rule r surjective_pairing)+ wenzelm@11838: apply (rule split_beta [THEN subst], rule q) wenzelm@11838: done wenzelm@11838: qed wenzelm@11838: wenzelm@11838: lemma splitD': "split R (a,b) c ==> R a b c" wenzelm@11838: by simp wenzelm@11838: wenzelm@11838: lemma mem_splitI: "z: c a b ==> z: split c (a, b)" wenzelm@11838: by simp wenzelm@11838: wenzelm@11838: lemma mem_splitI2: "!!p. [| !!a b. p = (a, b) ==> z: c a b |] ==> z: split c p" paulson@14208: by (simp only: split_tupled_all, simp) wenzelm@11838: wenzelm@18372: lemma mem_splitE: haftmann@37166: assumes major: "z \ split c p" haftmann@37166: and cases: "\x y. p = (x, y) \ z \ c x y \ Q" wenzelm@18372: shows Q haftmann@37591: by (rule major [unfolded prod_case_unfold] cases surjective_pairing)+ wenzelm@11838: wenzelm@11838: declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!] wenzelm@11838: declare mem_splitE [elim!] splitE' [elim!] splitE [elim!] wenzelm@11838: wenzelm@26340: ML {* wenzelm@11838: local (* filtering with exists_p_split is an essential optimization *) haftmann@37591: fun exists_p_split (Const (@{const_name prod_case},_) $ _ $ (Const (@{const_name Pair},_)$_$_)) = true wenzelm@11838: | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u wenzelm@11838: | exists_p_split (Abs (_, _, t)) = exists_p_split t wenzelm@11838: | exists_p_split _ = false; wenzelm@35364: val ss = HOL_basic_ss addsimps @{thms split_conv}; wenzelm@11838: in wenzelm@11838: val split_conv_tac = SUBGOAL (fn (t, i) => wenzelm@11838: if exists_p_split t then safe_full_simp_tac ss i else no_tac); wenzelm@11838: end; wenzelm@26340: *} wenzelm@26340: wenzelm@11838: (* This prevents applications of splitE for already splitted arguments leading wenzelm@11838: to quite time-consuming computations (in particular for nested tuples) *) wenzelm@26340: declaration {* fn _ => wenzelm@26340: Classical.map_cs (fn cs => cs addSbefore ("split_conv_tac", split_conv_tac)) wenzelm@16121: *} wenzelm@11838: blanchet@35828: lemma split_eta_SetCompr [simp,no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P" wenzelm@18372: by (rule ext) fast wenzelm@11838: blanchet@35828: lemma split_eta_SetCompr2 [simp,no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P" wenzelm@18372: by (rule ext) fast wenzelm@11838: wenzelm@11838: lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)" wenzelm@11838: -- {* Allows simplifications of nested splits in case of independent predicates. *} wenzelm@18372: by (rule ext) blast wenzelm@11838: nipkow@14337: (* Do NOT make this a simp rule as it nipkow@14337: a) only helps in special situations nipkow@14337: b) can lead to nontermination in the presence of split_def nipkow@14337: *) nipkow@14337: lemma split_comp_eq: paulson@20415: fixes f :: "'a => 'b => 'c" and g :: "'d => 'a" paulson@20415: shows "(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))" wenzelm@18372: by (rule ext) auto oheimb@14101: haftmann@26358: lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A" haftmann@26358: apply (rule_tac x = "(a, b)" in image_eqI) haftmann@26358: apply auto haftmann@26358: done haftmann@26358: wenzelm@11838: lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: (* wenzelm@11838: the following would be slightly more general, wenzelm@11838: but cannot be used as rewrite rule: wenzelm@11838: ### Cannot add premise as rewrite rule because it contains (type) unknowns: wenzelm@11838: ### ?y = .x wenzelm@11838: Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)" paulson@14208: by (rtac some_equality 1) paulson@14208: by ( Simp_tac 1) paulson@14208: by (split_all_tac 1) paulson@14208: by (Asm_full_simp_tac 1) wenzelm@11838: qed "The_split_eq"; wenzelm@11838: *) wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: Setup of internal @{text split_rule}. wenzelm@11838: *} wenzelm@11838: wenzelm@45607: lemmas prod_caseI = prod.cases [THEN iffD2] haftmann@24699: haftmann@24699: lemma prod_caseI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> prod_case c p" haftmann@37678: by (fact splitI2) haftmann@24699: haftmann@24699: lemma prod_caseI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> prod_case c p x" haftmann@37678: by (fact splitI2') haftmann@24699: haftmann@24699: lemma prod_caseE: "prod_case c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" haftmann@37678: by (fact splitE) haftmann@24699: haftmann@24699: lemma prod_caseE': "prod_case c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" haftmann@37678: by (fact splitE') haftmann@24699: haftmann@37678: declare prod_caseI [intro!] haftmann@24699: bulwahn@26143: lemma prod_case_beta: bulwahn@26143: "prod_case f p = f (fst p) (snd p)" haftmann@37591: by (fact split_beta) bulwahn@26143: haftmann@24699: lemma prod_cases3 [cases type]: haftmann@24699: obtains (fields) a b c where "y = (a, b, c)" haftmann@24699: by (cases y, case_tac b) blast haftmann@24699: haftmann@24699: lemma prod_induct3 [case_names fields, induct type]: haftmann@24699: "(!!a b c. P (a, b, c)) ==> P x" haftmann@24699: by (cases x) blast haftmann@24699: haftmann@24699: lemma prod_cases4 [cases type]: haftmann@24699: obtains (fields) a b c d where "y = (a, b, c, d)" haftmann@24699: by (cases y, case_tac c) blast haftmann@24699: haftmann@24699: lemma prod_induct4 [case_names fields, induct type]: haftmann@24699: "(!!a b c d. P (a, b, c, d)) ==> P x" haftmann@24699: by (cases x) blast haftmann@24699: haftmann@24699: lemma prod_cases5 [cases type]: haftmann@24699: obtains (fields) a b c d e where "y = (a, b, c, d, e)" haftmann@24699: by (cases y, case_tac d) blast haftmann@24699: haftmann@24699: lemma prod_induct5 [case_names fields, induct type]: haftmann@24699: "(!!a b c d e. P (a, b, c, d, e)) ==> P x" haftmann@24699: by (cases x) blast haftmann@24699: haftmann@24699: lemma prod_cases6 [cases type]: haftmann@24699: obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)" haftmann@24699: by (cases y, case_tac e) blast haftmann@24699: haftmann@24699: lemma prod_induct6 [case_names fields, induct type]: haftmann@24699: "(!!a b c d e f. P (a, b, c, d, e, f)) ==> P x" haftmann@24699: by (cases x) blast haftmann@24699: haftmann@24699: lemma prod_cases7 [cases type]: haftmann@24699: obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)" haftmann@24699: by (cases y, case_tac f) blast haftmann@24699: haftmann@24699: lemma prod_induct7 [case_names fields, induct type]: haftmann@24699: "(!!a b c d e f g. P (a, b, c, d, e, f, g)) ==> P x" haftmann@24699: by (cases x) blast haftmann@24699: haftmann@37166: lemma split_def: haftmann@37166: "split = (\c p. c (fst p) (snd p))" haftmann@37591: by (fact prod_case_unfold) haftmann@37166: haftmann@37166: definition internal_split :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where haftmann@37166: "internal_split == split" haftmann@37166: haftmann@37166: lemma internal_split_conv: "internal_split c (a, b) = c a b" haftmann@37166: by (simp only: internal_split_def split_conv) haftmann@37166: wenzelm@48891: ML_file "Tools/split_rule.ML" haftmann@37166: setup Split_Rule.setup haftmann@37166: haftmann@37166: hide_const internal_split haftmann@37166: haftmann@24699: haftmann@26358: subsubsection {* Derived operations *} haftmann@26358: haftmann@37387: definition curry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where haftmann@37387: "curry = (\c x y. c (x, y))" haftmann@37166: haftmann@37166: lemma curry_conv [simp, code]: "curry f a b = f (a, b)" haftmann@37166: by (simp add: curry_def) haftmann@37166: haftmann@37166: lemma curryI [intro!]: "f (a, b) \ curry f a b" haftmann@37166: by (simp add: curry_def) haftmann@37166: haftmann@37166: lemma curryD [dest!]: "curry f a b \ f (a, b)" haftmann@37166: by (simp add: curry_def) haftmann@37166: haftmann@37166: lemma curryE: "curry f a b \ (f (a, b) \ Q) \ Q" haftmann@37166: by (simp add: curry_def) haftmann@37166: haftmann@37166: lemma curry_split [simp]: "curry (split f) = f" haftmann@37166: by (simp add: curry_def split_def) haftmann@37166: haftmann@37166: lemma split_curry [simp]: "split (curry f) = f" haftmann@37166: by (simp add: curry_def split_def) haftmann@37166: haftmann@26358: text {* haftmann@26358: The composition-uncurry combinator. haftmann@26358: *} haftmann@26358: haftmann@37751: notation fcomp (infixl "\>" 60) haftmann@26358: haftmann@37751: definition scomp :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" (infixl "\\" 60) where haftmann@37751: "f \\ g = (\x. prod_case g (f x))" haftmann@26358: haftmann@37678: lemma scomp_unfold: "scomp = (\f g x. g (fst (f x)) (snd (f x)))" nipkow@39302: by (simp add: fun_eq_iff scomp_def prod_case_unfold) haftmann@37678: haftmann@37751: lemma scomp_apply [simp]: "(f \\ g) x = prod_case g (f x)" haftmann@37751: by (simp add: scomp_unfold prod_case_unfold) haftmann@26358: haftmann@37751: lemma Pair_scomp: "Pair x \\ f = f x" huffman@44921: by (simp add: fun_eq_iff) haftmann@26358: haftmann@37751: lemma scomp_Pair: "x \\ Pair = x" huffman@44921: by (simp add: fun_eq_iff) haftmann@26358: haftmann@37751: lemma scomp_scomp: "(f \\ g) \\ h = f \\ (\x. g x \\ h)" nipkow@39302: by (simp add: fun_eq_iff scomp_unfold) haftmann@26358: haftmann@37751: lemma scomp_fcomp: "(f \\ g) \> h = f \\ (\x. g x \> h)" nipkow@39302: by (simp add: fun_eq_iff scomp_unfold fcomp_def) haftmann@26358: haftmann@37751: lemma fcomp_scomp: "(f \> g) \\ h = f \> (g \\ h)" huffman@44921: by (simp add: fun_eq_iff scomp_unfold) haftmann@26358: haftmann@31202: code_const scomp haftmann@31202: (Eval infixl 3 "#->") haftmann@31202: haftmann@37751: no_notation fcomp (infixl "\>" 60) haftmann@37751: no_notation scomp (infixl "\\" 60) haftmann@26358: haftmann@26358: text {* haftmann@40607: @{term map_pair} --- action of the product functor upon krauss@36664: functions. haftmann@26358: *} haftmann@21195: haftmann@40607: definition map_pair :: "('a \ 'c) \ ('b \ 'd) \ 'a \ 'b \ 'c \ 'd" where haftmann@40607: "map_pair f g = (\(x, y). (f x, g y))" haftmann@26358: haftmann@40607: lemma map_pair_simp [simp, code]: haftmann@40607: "map_pair f g (a, b) = (f a, g b)" haftmann@40607: by (simp add: map_pair_def) haftmann@26358: haftmann@41505: enriched_type map_pair: map_pair huffman@44921: by (auto simp add: split_paired_all) nipkow@37278: haftmann@40607: lemma fst_map_pair [simp]: haftmann@40607: "fst (map_pair f g x) = f (fst x)" haftmann@40607: by (cases x) simp_all nipkow@37278: haftmann@40607: lemma snd_prod_fun [simp]: haftmann@40607: "snd (map_pair f g x) = g (snd x)" haftmann@40607: by (cases x) simp_all nipkow@37278: haftmann@40607: lemma fst_comp_map_pair [simp]: haftmann@40607: "fst \ map_pair f g = f \ fst" haftmann@40607: by (rule ext) simp_all nipkow@37278: haftmann@40607: lemma snd_comp_map_pair [simp]: haftmann@40607: "snd \ map_pair f g = g \ snd" haftmann@40607: by (rule ext) simp_all haftmann@26358: haftmann@40607: lemma map_pair_compose: haftmann@40607: "map_pair (f1 o f2) (g1 o g2) = (map_pair f1 g1 o map_pair f2 g2)" haftmann@40607: by (rule ext) (simp add: map_pair.compositionality comp_def) haftmann@26358: haftmann@40607: lemma map_pair_ident [simp]: haftmann@40607: "map_pair (%x. x) (%y. y) = (%z. z)" haftmann@40607: by (rule ext) (simp add: map_pair.identity) haftmann@40607: haftmann@40607: lemma map_pair_imageI [intro]: haftmann@40607: "(a, b) \ R \ (f a, g b) \ map_pair f g ` R" haftmann@40607: by (rule image_eqI) simp_all haftmann@21195: haftmann@26358: lemma prod_fun_imageE [elim!]: haftmann@40607: assumes major: "c \ map_pair f g ` R" haftmann@40607: and cases: "\x y. c = (f x, g y) \ (x, y) \ R \ P" haftmann@26358: shows P haftmann@26358: apply (rule major [THEN imageE]) haftmann@37166: apply (case_tac x) haftmann@26358: apply (rule cases) haftmann@40607: apply simp_all haftmann@26358: done haftmann@26358: haftmann@37166: definition apfst :: "('a \ 'c) \ 'a \ 'b \ 'c \ 'b" where haftmann@40607: "apfst f = map_pair f id" haftmann@26358: haftmann@37166: definition apsnd :: "('b \ 'c) \ 'a \ 'b \ 'a \ 'c" where haftmann@40607: "apsnd f = map_pair id f" haftmann@26358: haftmann@26358: lemma apfst_conv [simp, code]: haftmann@26358: "apfst f (x, y) = (f x, y)" haftmann@26358: by (simp add: apfst_def) haftmann@26358: hoelzl@33638: lemma apsnd_conv [simp, code]: haftmann@26358: "apsnd f (x, y) = (x, f y)" haftmann@26358: by (simp add: apsnd_def) haftmann@21195: haftmann@33594: lemma fst_apfst [simp]: haftmann@33594: "fst (apfst f x) = f (fst x)" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma fst_apsnd [simp]: haftmann@33594: "fst (apsnd f x) = fst x" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma snd_apfst [simp]: haftmann@33594: "snd (apfst f x) = snd x" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma snd_apsnd [simp]: haftmann@33594: "snd (apsnd f x) = f (snd x)" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma apfst_compose: haftmann@33594: "apfst f (apfst g x) = apfst (f \ g) x" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma apsnd_compose: haftmann@33594: "apsnd f (apsnd g x) = apsnd (f \ g) x" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma apfst_apsnd [simp]: haftmann@33594: "apfst f (apsnd g x) = (f (fst x), g (snd x))" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma apsnd_apfst [simp]: haftmann@33594: "apsnd f (apfst g x) = (g (fst x), f (snd x))" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma apfst_id [simp] : haftmann@33594: "apfst id = id" nipkow@39302: by (simp add: fun_eq_iff) haftmann@33594: haftmann@33594: lemma apsnd_id [simp] : haftmann@33594: "apsnd id = id" nipkow@39302: by (simp add: fun_eq_iff) haftmann@33594: haftmann@33594: lemma apfst_eq_conv [simp]: haftmann@33594: "apfst f x = apfst g x \ f (fst x) = g (fst x)" haftmann@33594: by (cases x) simp haftmann@33594: haftmann@33594: lemma apsnd_eq_conv [simp]: haftmann@33594: "apsnd f x = apsnd g x \ f (snd x) = g (snd x)" haftmann@33594: by (cases x) simp haftmann@33594: hoelzl@33638: lemma apsnd_apfst_commute: hoelzl@33638: "apsnd f (apfst g p) = apfst g (apsnd f p)" hoelzl@33638: by simp haftmann@21195: haftmann@26358: text {* haftmann@26358: Disjoint union of a family of sets -- Sigma. haftmann@26358: *} haftmann@26358: haftmann@45986: definition Sigma :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" where haftmann@26358: Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}" haftmann@26358: haftmann@26358: abbreviation haftmann@45986: Times :: "'a set \ 'b set \ ('a \ 'b) set" haftmann@26358: (infixr "<*>" 80) where haftmann@26358: "A <*> B == Sigma A (%_. B)" haftmann@26358: haftmann@26358: notation (xsymbols) haftmann@26358: Times (infixr "\" 80) berghofe@15394: haftmann@26358: notation (HTML output) haftmann@26358: Times (infixr "\" 80) haftmann@26358: nipkow@45662: hide_const (open) Times nipkow@45662: haftmann@26358: syntax wenzelm@35115: "_Sigma" :: "[pttrn, 'a set, 'b set] => ('a * 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10) haftmann@26358: translations wenzelm@35115: "SIGMA x:A. B" == "CONST Sigma A (%x. B)" haftmann@26358: haftmann@26358: lemma SigmaI [intro!]: "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" haftmann@26358: by (unfold Sigma_def) blast haftmann@26358: haftmann@26358: lemma SigmaE [elim!]: haftmann@26358: "[| c: Sigma A B; haftmann@26358: !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P haftmann@26358: |] ==> P" haftmann@26358: -- {* The general elimination rule. *} haftmann@26358: by (unfold Sigma_def) blast haftmann@20588: haftmann@26358: text {* haftmann@26358: Elimination of @{term "(a, b) : A \ B"} -- introduces no haftmann@26358: eigenvariables. haftmann@26358: *} haftmann@26358: haftmann@26358: lemma SigmaD1: "(a, b) : Sigma A B ==> a : A" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma SigmaE2: haftmann@26358: "[| (a, b) : Sigma A B; haftmann@26358: [| a:A; b:B(a) |] ==> P haftmann@26358: |] ==> P" haftmann@26358: by blast haftmann@20588: haftmann@26358: lemma Sigma_cong: haftmann@26358: "\A = B; !!x. x \ B \ C x = D x\ haftmann@26358: \ (SIGMA x: A. C x) = (SIGMA x: B. D x)" haftmann@26358: by auto haftmann@26358: haftmann@26358: lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Sigma_empty1 [simp]: "Sigma {} B = {}" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Sigma_empty2 [simp]: "A <*> {} = {}" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV" haftmann@26358: by auto haftmann@21908: haftmann@26358: lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)" haftmann@26358: by auto haftmann@26358: haftmann@26358: lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV" haftmann@26358: by auto haftmann@26358: haftmann@26358: lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)" haftmann@26358: by (blast elim: equalityE) haftmann@20588: haftmann@26358: lemma SetCompr_Sigma_eq: haftmann@26358: "Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma UN_Times_distrib: haftmann@26358: "(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)" haftmann@26358: -- {* Suggested by Pierre Chartier *} haftmann@26358: by blast haftmann@26358: blanchet@47740: lemma split_paired_Ball_Sigma [simp, no_atp]: haftmann@26358: "(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))" haftmann@26358: by blast haftmann@26358: blanchet@47740: lemma split_paired_Bex_Sigma [simp, no_atp]: haftmann@26358: "(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))" haftmann@26358: by blast haftmann@21908: haftmann@26358: lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))" haftmann@26358: by blast haftmann@26358: haftmann@26358: lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))" haftmann@26358: by blast haftmann@21908: haftmann@26358: lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)" haftmann@26358: by blast haftmann@26358: haftmann@26358: text {* haftmann@26358: Non-dependent versions are needed to avoid the need for higher-order haftmann@26358: matching, especially when the rules are re-oriented. haftmann@26358: *} haftmann@21908: haftmann@26358: lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)" nipkow@28719: by blast haftmann@26358: haftmann@26358: lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)" nipkow@28719: by blast haftmann@26358: haftmann@26358: lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)" nipkow@28719: by blast haftmann@26358: hoelzl@36622: lemma Times_empty[simp]: "A \ B = {} \ A = {} \ B = {}" hoelzl@36622: by auto hoelzl@36622: hoelzl@36622: lemma fst_image_times[simp]: "fst ` (A \ B) = (if B = {} then {} else A)" huffman@44921: by force hoelzl@36622: hoelzl@36622: lemma snd_image_times[simp]: "snd ` (A \ B) = (if A = {} then {} else B)" huffman@44921: by force hoelzl@36622: nipkow@28719: lemma insert_times_insert[simp]: nipkow@28719: "insert a A \ insert b B = nipkow@28719: insert (a,b) (A \ insert b B \ insert a A \ B)" nipkow@28719: by blast haftmann@26358: paulson@33271: lemma vimage_Times: "f -` (A \ B) = ((fst \ f) -` A) \ ((snd \ f) -` B)" wenzelm@47988: apply auto wenzelm@47988: apply (case_tac "f x") wenzelm@47988: apply auto wenzelm@47988: done paulson@33271: haftmann@35822: lemma swap_inj_on: hoelzl@36622: "inj_on (\(i, j). (j, i)) A" hoelzl@36622: by (auto intro!: inj_onI) haftmann@35822: haftmann@35822: lemma swap_product: haftmann@35822: "(%(i, j). (j, i)) ` (A \ B) = B \ A" haftmann@35822: by (simp add: split_def image_def) blast haftmann@35822: hoelzl@36622: lemma image_split_eq_Sigma: hoelzl@36622: "(\x. (f x, g x)) ` A = Sigma (f ` A) (\x. g ` (f -` {x} \ A))" haftmann@46128: proof (safe intro!: imageI) hoelzl@36622: fix a b assume *: "a \ A" "b \ A" and eq: "f a = f b" hoelzl@36622: show "(f b, g a) \ (\x. (f x, g x)) ` A" hoelzl@36622: using * eq[symmetric] by auto hoelzl@36622: qed simp_all haftmann@35822: haftmann@46128: definition product :: "'a set \ 'b set \ ('a \ 'b) set" where haftmann@46128: [code_abbrev]: "product A B = A \ B" haftmann@46128: haftmann@46128: hide_const (open) product haftmann@46128: haftmann@46128: lemma member_product: haftmann@46128: "x \ Product_Type.product A B \ x \ A \ B" haftmann@46128: by (simp add: product_def) haftmann@46128: haftmann@40607: text {* The following @{const map_pair} lemmas are due to Joachim Breitner: *} haftmann@40607: haftmann@40607: lemma map_pair_inj_on: haftmann@40607: assumes "inj_on f A" and "inj_on g B" haftmann@40607: shows "inj_on (map_pair f g) (A \ B)" haftmann@40607: proof (rule inj_onI) haftmann@40607: fix x :: "'a \ 'c" and y :: "'a \ 'c" haftmann@40607: assume "x \ A \ B" hence "fst x \ A" and "snd x \ B" by auto haftmann@40607: assume "y \ A \ B" hence "fst y \ A" and "snd y \ B" by auto haftmann@40607: assume "map_pair f g x = map_pair f g y" haftmann@40607: hence "fst (map_pair f g x) = fst (map_pair f g y)" by (auto) haftmann@40607: hence "f (fst x) = f (fst y)" by (cases x,cases y,auto) haftmann@40607: with `inj_on f A` and `fst x \ A` and `fst y \ A` haftmann@40607: have "fst x = fst y" by (auto dest:dest:inj_onD) haftmann@40607: moreover from `map_pair f g x = map_pair f g y` haftmann@40607: have "snd (map_pair f g x) = snd (map_pair f g y)" by (auto) haftmann@40607: hence "g (snd x) = g (snd y)" by (cases x,cases y,auto) haftmann@40607: with `inj_on g B` and `snd x \ B` and `snd y \ B` haftmann@40607: have "snd x = snd y" by (auto dest:dest:inj_onD) haftmann@40607: ultimately show "x = y" by(rule prod_eqI) haftmann@40607: qed haftmann@40607: haftmann@40607: lemma map_pair_surj: hoelzl@40702: fixes f :: "'a \ 'b" and g :: "'c \ 'd" haftmann@40607: assumes "surj f" and "surj g" haftmann@40607: shows "surj (map_pair f g)" haftmann@40607: unfolding surj_def haftmann@40607: proof haftmann@40607: fix y :: "'b \ 'd" haftmann@40607: from `surj f` obtain a where "fst y = f a" by (auto elim:surjE) haftmann@40607: moreover haftmann@40607: from `surj g` obtain b where "snd y = g b" by (auto elim:surjE) haftmann@40607: ultimately have "(fst y, snd y) = map_pair f g (a,b)" by auto haftmann@40607: thus "\x. y = map_pair f g x" by auto haftmann@40607: qed haftmann@40607: haftmann@40607: lemma map_pair_surj_on: haftmann@40607: assumes "f ` A = A'" and "g ` B = B'" haftmann@40607: shows "map_pair f g ` (A \ B) = A' \ B'" haftmann@40607: unfolding image_def haftmann@40607: proof(rule set_eqI,rule iffI) haftmann@40607: fix x :: "'a \ 'c" haftmann@40607: assume "x \ {y\'a \ 'c. \x\'b \ 'd\A \ B. y = map_pair f g x}" haftmann@40607: then obtain y where "y \ A \ B" and "x = map_pair f g y" by blast haftmann@40607: from `image f A = A'` and `y \ A \ B` have "f (fst y) \ A'" by auto haftmann@40607: moreover from `image g B = B'` and `y \ A \ B` have "g (snd y) \ B'" by auto haftmann@40607: ultimately have "(f (fst y), g (snd y)) \ (A' \ B')" by auto haftmann@40607: with `x = map_pair f g y` show "x \ A' \ B'" by (cases y, auto) haftmann@40607: next haftmann@40607: fix x :: "'a \ 'c" haftmann@40607: assume "x \ A' \ B'" hence "fst x \ A'" and "snd x \ B'" by auto haftmann@40607: from `image f A = A'` and `fst x \ A'` have "fst x \ image f A" by auto haftmann@40607: then obtain a where "a \ A" and "fst x = f a" by (rule imageE) haftmann@40607: moreover from `image g B = B'` and `snd x \ B'` haftmann@40607: obtain b where "b \ B" and "snd x = g b" by auto haftmann@40607: ultimately have "(fst x, snd x) = map_pair f g (a,b)" by auto haftmann@40607: moreover from `a \ A` and `b \ B` have "(a , b) \ A \ B" by auto haftmann@40607: ultimately have "\y \ A \ B. x = map_pair f g y" by auto haftmann@40607: thus "x \ {x. \y \ A \ B. x = map_pair f g y}" by auto haftmann@40607: qed haftmann@40607: haftmann@21908: bulwahn@49764: subsection {* Simproc for rewriting a set comprehension into a pointfree expression *} bulwahn@49764: bulwahn@49764: ML_file "Tools/set_comprehension_pointfree.ML" bulwahn@49764: bulwahn@49764: setup {* bulwahn@49764: Code_Preproc.map_pre (fn ss => ss addsimprocs bulwahn@49764: [Raw_Simplifier.make_simproc {name = "set comprehension", lhss = [@{cpat "Collect ?P"}], bulwahn@49764: proc = K Set_Comprehension_Pointfree.code_simproc, identifier = []}]) bulwahn@49764: *} bulwahn@49764: bulwahn@49764: haftmann@37166: subsection {* Inductively defined sets *} berghofe@15394: wenzelm@48891: ML_file "Tools/inductive_set.ML" haftmann@31723: setup Inductive_Set.setup haftmann@24699: haftmann@37166: haftmann@37166: subsection {* Legacy theorem bindings and duplicates *} haftmann@37166: haftmann@37166: lemma PairE: haftmann@37166: obtains x y where "p = (x, y)" haftmann@37166: by (fact prod.exhaust) haftmann@37166: haftmann@37166: lemma Pair_inject: haftmann@37166: assumes "(a, b) = (a', b')" haftmann@37166: and "a = a' ==> b = b' ==> R" haftmann@37166: shows R haftmann@37166: using assms by simp haftmann@37166: haftmann@37166: lemmas Pair_eq = prod.inject haftmann@37166: haftmann@37166: lemmas split = split_conv -- {* for backwards compatibility *} haftmann@37166: huffman@44066: lemmas Pair_fst_snd_eq = prod_eq_iff huffman@44066: huffman@45204: hide_const (open) prod huffman@45204: nipkow@10213: end