nipkow@10213: (* Title: HOL/Product_Type.thy nipkow@10213: ID: $Id$ 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@21331: imports Typedef Fun Code_Generator haftmann@16417: uses ("Tools/split_rule.ML") nipkow@15131: begin wenzelm@11838: wenzelm@11838: subsection {* Unit *} wenzelm@11838: wenzelm@11838: typedef unit = "{True}" wenzelm@11838: proof haftmann@20588: show "True : ?unit" .. wenzelm@11838: qed wenzelm@11838: wenzelm@11838: constdefs wenzelm@11838: Unity :: unit ("'(')") wenzelm@11838: "() == Abs_unit True" wenzelm@11838: wenzelm@11838: lemma unit_eq: "u = ()" wenzelm@11838: by (induct u) (simp add: unit_def 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@11838: ML_setup {* wenzelm@13462: val unit_eq_proc = wenzelm@13462: let val unit_meta_eq = mk_meta_eq (thm "unit_eq") in haftmann@22349: Simplifier.simproc (the_context ()) "unit_eq" ["x::unit"] skalberg@15531: (fn _ => fn _ => fn t => if HOLogic.is_unit t then NONE else SOME unit_meta_eq) wenzelm@13462: end; wenzelm@11838: wenzelm@11838: Addsimprocs [unit_eq_proc]; wenzelm@11838: *} wenzelm@11838: 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: lemma unit_induct [induct type: unit]: "P () ==> P x" wenzelm@11838: by simp wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: This rewrite counters the effect of @{text unit_eq_proc} 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: wenzelm@11838: lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f" wenzelm@11838: by (rule ext) simp nipkow@10213: nipkow@10213: wenzelm@11838: subsection {* Pairs *} nipkow@10213: wenzelm@11777: subsubsection {* Type definition *} nipkow@10213: nipkow@10213: constdefs oheimb@11025: Pair_Rep :: "['a, 'b] => ['a, 'b] => bool" wenzelm@11032: "Pair_Rep == (%a b. %x y. x=a & y=b)" nipkow@10213: nipkow@10213: global nipkow@10213: nipkow@10213: typedef (Prod) wenzelm@11838: ('a, 'b) "*" (infixr 20) wenzelm@11032: = "{f. EX a b. f = Pair_Rep (a::'a) (b::'b)}" oheimb@11025: proof oheimb@11025: fix a b show "Pair_Rep a b : ?Prod" oheimb@11025: by blast oheimb@11025: qed nipkow@10213: wenzelm@12114: syntax (xsymbols) schirmer@15422: "*" :: "[type, type] => type" ("(_ \/ _)" [21, 20] 20) nipkow@10213: syntax (HTML output) schirmer@15422: "*" :: "[type, type] => type" ("(_ \/ _)" [21, 20] 20) nipkow@10213: wenzelm@11777: local nipkow@10213: wenzelm@11777: wenzelm@19535: subsubsection {* Definitions *} wenzelm@11777: wenzelm@11777: global nipkow@10213: nipkow@10213: consts oheimb@11025: fst :: "'a * 'b => 'a" oheimb@11025: snd :: "'a * 'b => 'b" oheimb@11025: split :: "[['a, 'b] => 'c, 'a * 'b] => 'c" skalberg@14189: curry :: "['a * 'b => 'c, 'a, 'b] => 'c" oheimb@11025: prod_fun :: "['a => 'b, 'c => 'd, 'a * 'c] => 'b * 'd" oheimb@11025: Pair :: "['a, 'b] => 'a * 'b" oheimb@11025: Sigma :: "['a set, 'a => 'b set] => ('a * 'b) set" nipkow@10213: wenzelm@11777: local nipkow@10213: wenzelm@19535: defs wenzelm@19535: Pair_def: "Pair a b == Abs_Prod (Pair_Rep a b)" wenzelm@19535: fst_def: "fst p == THE a. EX b. p = Pair a b" wenzelm@19535: snd_def: "snd p == THE b. EX a. p = Pair a b" wenzelm@19535: split_def: "split == (%c p. c (fst p) (snd p))" wenzelm@19535: curry_def: "curry == (%c x y. c (Pair x y))" wenzelm@19535: prod_fun_def: "prod_fun f g == split (%x y. Pair (f x) (g y))" wenzelm@19535: Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}" wenzelm@19535: wenzelm@19535: abbreviation wenzelm@21404: Times :: "['a set, 'b set] => ('a * 'b) set" wenzelm@21404: (infixr "<*>" 80) where wenzelm@19535: "A <*> B == Sigma A (%_. B)" wenzelm@19535: wenzelm@21210: notation (xsymbols) wenzelm@19656: Times (infixr "\" 80) wenzelm@19535: wenzelm@21210: notation (HTML output) wenzelm@19656: Times (infixr "\" 80) wenzelm@19535: wenzelm@19535: wenzelm@19535: subsubsection {* Concrete syntax *} wenzelm@19535: wenzelm@11777: text {* wenzelm@11777: Patterns -- extends pre-defined type @{typ pttrn} used in wenzelm@11777: abstractions. wenzelm@11777: *} nipkow@10213: nipkow@10213: nonterminals nipkow@10213: tuple_args 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" ("_,/ _") wenzelm@22439: "@Sigma" ::"[pttrn, 'a set, 'b set] => ('a * 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10) nipkow@10213: nipkow@10213: translations nipkow@10213: "(x, y)" == "Pair x y" nipkow@10213: "_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))" nipkow@10213: "%(x,y,zs).b" == "split(%x (y,zs).b)" nipkow@10213: "%(x,y).b" == "split(%x y. b)" nipkow@10213: "_abs (Pair x y) t" => "%(x,y).t" nipkow@10213: (* The last rule accommodates tuples in `case C ... (x,y) ... => ...' nipkow@10213: The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *) wenzelm@19535: "SIGMA x:A. B" == "Sigma A (%x. B)" nipkow@10213: schirmer@14359: (* reconstructs pattern from (nested) splits, avoiding eta-contraction of body*) schirmer@14359: (* works best with enclosing "let", if "let" does not avoid eta-contraction *) schirmer@14359: print_translation {* schirmer@14359: let fun split_tr' [Abs (x,T,t as (Abs abs))] = schirmer@14359: (* split (%x y. t) => %(x,y) t *) schirmer@14359: let val (y,t') = atomic_abs_tr' abs; schirmer@14359: val (x',t'') = atomic_abs_tr' (x,T,t'); schirmer@14359: schirmer@14359: in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end schirmer@14359: | split_tr' [Abs (x,T,(s as Const ("split",_)$t))] = schirmer@14359: (* split (%x. (split (%y z. t))) => %(x,y,z). t *) schirmer@14359: let val (Const ("_abs",_)$(Const ("_pattern",_)$y$z)$t') = split_tr' [t]; schirmer@14359: val (x',t'') = atomic_abs_tr' (x,T,t'); schirmer@14359: in Syntax.const "_abs"$ schirmer@14359: (Syntax.const "_pattern"$x'$(Syntax.const "_patterns"$y$z))$t'' end schirmer@14359: | split_tr' [Const ("split",_)$t] = schirmer@14359: (* split (split (%x y z. t)) => %((x,y),z). t *) schirmer@14359: split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *) schirmer@14359: | split_tr' [Const ("_abs",_)$x_y$(Abs abs)] = schirmer@14359: (* split (%pttrn z. t) => %(pttrn,z). t *) schirmer@14359: let val (z,t) = atomic_abs_tr' abs; schirmer@14359: in Syntax.const "_abs" $ (Syntax.const "_pattern" $x_y$z) $ t end schirmer@14359: | split_tr' _ = raise Match; schirmer@14359: in [("split", split_tr')] schirmer@14359: 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 schirmer@15422: fun split_guess_names_tr' _ T [Abs (x,_,Abs _)] = raise Match schirmer@15422: | split_guess_names_tr' _ T [Abs (x,xT,t)] = schirmer@15422: (case (head_of t) of schirmer@15422: Const ("split",_) => raise Match schirmer@15422: | _ => let schirmer@15422: val (_::yT::_) = binder_types (domain_type T) handle Bind => raise Match; schirmer@15422: val (y,t') = atomic_abs_tr' ("y",yT,(incr_boundvars 1 t)$Bound 0); schirmer@15422: val (x',t'') = atomic_abs_tr' (x,xT,t'); schirmer@15422: in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end) schirmer@15422: | split_guess_names_tr' _ T [t] = schirmer@15422: (case (head_of t) of schirmer@15422: Const ("split",_) => raise Match schirmer@15422: | _ => let schirmer@15422: val (xT::yT::_) = binder_types (domain_type T) handle Bind => raise Match; schirmer@15422: val (y,t') = schirmer@15422: atomic_abs_tr' ("y",yT,(incr_boundvars 2 t)$Bound 1$Bound 0); schirmer@15422: val (x',t'') = atomic_abs_tr' ("x",xT,t'); schirmer@15422: in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end) schirmer@15422: | split_guess_names_tr' _ _ _ = raise Match; schirmer@15422: in [("split", split_guess_names_tr')] schirmer@15422: end schirmer@15422: *} schirmer@15422: nipkow@10213: wenzelm@11966: subsubsection {* Lemmas and proof tool setup *} wenzelm@11838: wenzelm@11838: lemma ProdI: "Pair_Rep a b : Prod" wenzelm@19535: unfolding Prod_def by blast wenzelm@11838: wenzelm@11838: lemma Pair_Rep_inject: "Pair_Rep a b = Pair_Rep a' b' ==> a = a' & b = b'" wenzelm@11838: apply (unfold Pair_Rep_def) paulson@14208: apply (drule fun_cong [THEN fun_cong], blast) wenzelm@11838: done nipkow@10213: wenzelm@11838: lemma inj_on_Abs_Prod: "inj_on Abs_Prod Prod" wenzelm@11838: apply (rule inj_on_inverseI) wenzelm@11838: apply (erule Abs_Prod_inverse) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma Pair_inject: wenzelm@18372: assumes "(a, b) = (a', b')" wenzelm@18372: and "a = a' ==> b = b' ==> R" wenzelm@18372: shows R wenzelm@18372: apply (insert prems [unfolded Pair_def]) wenzelm@18372: apply (rule inj_on_Abs_Prod [THEN inj_onD, THEN Pair_Rep_inject, THEN conjE]) wenzelm@18372: apply (assumption | rule ProdI)+ wenzelm@18372: done nipkow@10213: wenzelm@11838: lemma Pair_eq [iff]: "((a, b) = (a', b')) = (a = a' & b = b')" wenzelm@11838: by (blast elim!: Pair_inject) wenzelm@11838: wenzelm@11838: lemma fst_conv [simp]: "fst (a, b) = a" wenzelm@19535: unfolding fst_def by blast wenzelm@11838: wenzelm@11838: lemma snd_conv [simp]: "snd (a, b) = b" wenzelm@19535: unfolding snd_def by blast oheimb@11025: 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: wenzelm@11838: lemma PairE_lemma: "EX x y. p = (x, y)" wenzelm@11838: apply (unfold Pair_def) wenzelm@11838: apply (rule Rep_Prod [unfolded Prod_def, THEN CollectE]) wenzelm@11838: apply (erule exE, erule exE, rule exI, rule exI) wenzelm@11838: apply (rule Rep_Prod_inverse [symmetric, THEN trans]) wenzelm@11838: apply (erule arg_cong) wenzelm@11838: done wenzelm@11032: wenzelm@11838: lemma PairE [cases type: *]: "(!!x y. p = (x, y) ==> Q) ==> Q" wenzelm@19535: using PairE_lemma [of p] by blast wenzelm@11838: wenzelm@16121: ML {* wenzelm@11838: local val PairE = thm "PairE" in wenzelm@11838: fun pair_tac s = wenzelm@11838: EVERY' [res_inst_tac [("p", s)] PairE, hyp_subst_tac, K prune_params_tac]; wenzelm@11838: end; wenzelm@11838: *} wenzelm@11032: wenzelm@11838: lemma surjective_pairing: "p = (fst p, snd p)" wenzelm@11838: -- {* Do not add as rewrite rule: invalidates some proofs in IMP *} wenzelm@11838: by (cases p) simp wenzelm@11838: paulson@17085: lemmas pair_collapse = surjective_pairing [symmetric] paulson@17085: declare pair_collapse [simp] oheimb@11025: wenzelm@11838: lemma surj_pair [simp]: "EX x y. z = (x, y)" wenzelm@11838: apply (rule exI) wenzelm@11838: apply (rule exI) wenzelm@11838: apply (rule surjective_pairing) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma split_paired_all: "(!!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: lemmas split_tupled_all = split_paired_all unit_all_eq2 wenzelm@11838: 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: wenzelm@16121: ML_setup {* 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@16121: addsimps [thm "split_paired_all", thm "unit_all_eq2", thm "unit_abs_eta_conv"] wenzelm@11838: addsimprocs [unit_eq_proc]; 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@11838: if exists_paired_all (#prop (Thm.rep_thm th)) then full_simplify ss th else th; wenzelm@11838: end; wenzelm@11838: wenzelm@17875: change_claset (fn cs => cs addSbefore ("split_all_tac", split_all_tac)); wenzelm@16121: *} wenzelm@11838: wenzelm@11838: lemma split_paired_All [simp]: "(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: skalberg@14189: lemma curry_split [simp]: "curry (split f) = f" skalberg@14189: by (simp add: curry_def split_def) skalberg@14189: skalberg@14189: lemma split_curry [simp]: "split (curry f) = f" skalberg@14189: by (simp add: curry_def split_def) skalberg@14189: skalberg@14189: lemma curryI [intro!]: "f (a,b) ==> curry f a b" skalberg@14189: by (simp add: curry_def) skalberg@14189: skalberg@14190: lemma curryD [dest!]: "curry f a b ==> f (a,b)" skalberg@14189: by (simp add: curry_def) skalberg@14189: skalberg@14190: lemma curryE: "[| curry f a b ; f (a,b) ==> Q |] ==> Q" skalberg@14189: by (simp add: curry_def) skalberg@14189: skalberg@14189: lemma curry_conv [simp]: "curry f a b = f (a,b)" skalberg@14189: by (simp add: curry_def) skalberg@14189: wenzelm@11838: lemma prod_induct [induct type: *]: "!!x. (!!a b. P (a, b)) ==> P x" wenzelm@11838: by fast wenzelm@11838: wenzelm@11838: lemma split_paired_Ex [simp]: "(EX x. P x) = (EX a b. P (a, b))" wenzelm@11838: by fast wenzelm@11838: wenzelm@11838: lemma split_conv [simp]: "split c (a, b) = c a b" wenzelm@11838: by (simp add: split_def) wenzelm@11838: wenzelm@11838: lemmas split = split_conv -- {* for backwards compatibility *} wenzelm@11838: wenzelm@11838: lemmas splitI = split_conv [THEN iffD2, standard] wenzelm@11838: lemmas splitD = split_conv [THEN iffD1, standard] wenzelm@11838: wenzelm@11838: lemma split_Pair_apply: "split (%x y. f (x, y)) = f" wenzelm@11838: -- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *} wenzelm@11838: apply (rule ext) paulson@14208: apply (tactic {* pair_tac "x" 1 *}, simp) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma split_paired_The: "(THE x. P x) = (THE (a, b). P (a, b))" wenzelm@11838: -- {* Can't be added to simpset: loops! *} wenzelm@11838: by (simp add: split_Pair_apply) wenzelm@11838: wenzelm@11838: lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))" wenzelm@11838: by (simp add: split_def) wenzelm@11838: wenzelm@11838: lemma Pair_fst_snd_eq: "!!s t. (s = t) = (fst s = fst t & snd s = snd t)" paulson@14208: by (simp only: split_tupled_all, simp) wenzelm@11838: wenzelm@11838: lemma prod_eqI [intro?]: "fst p = fst q ==> snd p = snd q ==> p = q" wenzelm@11838: by (simp add: Pair_fst_snd_eq) wenzelm@11838: wenzelm@11838: lemma split_weak_cong: "p = q ==> split c p = split c q" wenzelm@11838: -- {* Prevents simplification of @{term c}: much faster *} wenzelm@11838: by (erule arg_cong) wenzelm@11838: wenzelm@11838: lemma split_eta: "(%(x, y). f (x, y)) = f" wenzelm@11838: apply (rule ext) wenzelm@11838: apply (simp only: split_tupled_all) wenzelm@11838: apply (rule split_conv) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g" wenzelm@11838: 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 wenzelm@11838: split_beta}. *} wenzelm@11838: wenzelm@11838: ML_setup {* wenzelm@11838: wenzelm@11838: local wenzelm@18328: val cond_split_eta_ss = HOL_basic_ss addsimps [thm "cond_split_eta"] wenzelm@11838: fun Pair_pat k 0 (Bound m) = (m = k) wenzelm@11838: | Pair_pat k i (Const ("Pair", _) $ Bound m $ t) = i > 0 andalso wenzelm@11838: m = k+i andalso Pair_pat k (i-1) t wenzelm@11838: | Pair_pat _ _ _ = false; wenzelm@11838: fun no_args k i (Abs (_, _, t)) = no_args (k+1) i t wenzelm@11838: | no_args k i (t $ u) = no_args k i t andalso no_args k i u wenzelm@11838: | no_args k i (Bound m) = m < k orelse m > k+i wenzelm@11838: | no_args _ _ _ = true; skalberg@15531: fun split_pat tp i (Abs (_,_,t)) = if tp 0 i t then SOME (i,t) else NONE wenzelm@11838: | split_pat tp i (Const ("split", _) $ Abs (_, _, t)) = split_pat tp (i+1) t skalberg@15531: | split_pat tp i _ = NONE; wenzelm@20044: fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] [] wenzelm@13480: (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@11838: fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k+1) i t wenzelm@11838: | beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse wenzelm@11838: (beta_term_pat k i t andalso beta_term_pat k i u) wenzelm@11838: | beta_term_pat k i t = no_args k i t; wenzelm@11838: fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg wenzelm@11838: | 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@11838: | subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg wenzelm@11838: else (subst arg k i t $ subst arg k i u) wenzelm@11838: | subst arg k i t = t; wenzelm@20044: fun beta_proc ss (s as Const ("split", _) $ Abs (_, _, t) $ arg) = wenzelm@11838: (case split_pat beta_term_pat 1 t of wenzelm@20044: SOME (i,f) => SOME (metaeq ss s (subst arg 0 i f)) skalberg@15531: | NONE => NONE) wenzelm@20044: | beta_proc _ _ = NONE; wenzelm@20044: fun eta_proc ss (s as Const ("split", _) $ Abs (_, _, t)) = wenzelm@11838: (case split_pat eta_term_pat 1 t of wenzelm@20044: SOME (_,ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end)) skalberg@15531: | NONE => NONE) wenzelm@20044: | eta_proc _ _ = NONE; wenzelm@11838: in wenzelm@13462: val split_beta_proc = Simplifier.simproc (Theory.sign_of (the_context ())) wenzelm@20044: "split_beta" ["split f z"] (K beta_proc); wenzelm@13462: val split_eta_proc = Simplifier.simproc (Theory.sign_of (the_context ())) wenzelm@20044: "split_eta" ["split f"] (K eta_proc); wenzelm@11838: end; wenzelm@11838: wenzelm@11838: Addsimprocs [split_beta_proc, split_eta_proc]; wenzelm@11838: *} wenzelm@11838: wenzelm@11838: lemma split_beta: "(%(x, y). P x y) z = P (fst z) (snd z)" wenzelm@11838: by (subst surjective_pairing, rule split_conv) wenzelm@11838: wenzelm@11838: lemma split_split: "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: wenzelm@11838: lemma split_split_asm: "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: 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 wenzelm@11838: call @{text simp} using @{thm [source] 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" wenzelm@11838: by (induct p) (auto simp add: split_def) wenzelm@11838: wenzelm@11838: lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" wenzelm@11838: by (induct p) (auto simp add: split_def) 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: wenzelm@18372: assumes major: "z: split c p" wenzelm@18372: and cases: "!!x y. [| p = (x,y); z: c x y |] ==> Q" wenzelm@18372: shows Q wenzelm@18372: by (rule major [unfolded split_def] 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@16121: ML_setup {* wenzelm@11838: local (* filtering with exists_p_split is an essential optimization *) wenzelm@16121: fun exists_p_split (Const ("split",_) $ _ $ (Const ("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@16121: val ss = HOL_basic_ss addsimps [thm "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@11838: (* This prevents applications of splitE for already splitted arguments leading wenzelm@11838: to quite time-consuming computations (in particular for nested tuples) *) wenzelm@17875: change_claset (fn cs => cs addSbefore ("split_conv_tac", split_conv_tac)); wenzelm@16121: *} wenzelm@11838: wenzelm@11838: lemma split_eta_SetCompr [simp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P" wenzelm@18372: by (rule ext) fast wenzelm@11838: wenzelm@11838: lemma split_eta_SetCompr2 [simp]: "(%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: 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: lemma injective_fst_snd: "!!x y. [|fst x = fst y; snd x = snd y|] ==> x = y" wenzelm@11838: by auto wenzelm@11838: wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: \bigskip @{term prod_fun} --- action of the product functor upon wenzelm@11838: functions. wenzelm@11838: *} wenzelm@11838: wenzelm@11838: lemma prod_fun [simp]: "prod_fun f g (a, b) = (f a, g b)" wenzelm@11838: by (simp add: prod_fun_def) wenzelm@11838: wenzelm@11838: lemma prod_fun_compose: "prod_fun (f1 o f2) (g1 o g2) = (prod_fun f1 g1 o prod_fun f2 g2)" wenzelm@11838: apply (rule ext) paulson@14208: apply (tactic {* pair_tac "x" 1 *}, simp) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma prod_fun_ident [simp]: "prod_fun (%x. x) (%y. y) = (%z. z)" wenzelm@11838: apply (rule ext) paulson@14208: apply (tactic {* pair_tac "z" 1 *}, simp) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma prod_fun_imageI [intro]: "(a, b) : r ==> (f a, g b) : prod_fun f g ` r" wenzelm@11838: apply (rule image_eqI) paulson@14208: apply (rule prod_fun [symmetric], assumption) wenzelm@11838: done wenzelm@11838: wenzelm@11838: lemma prod_fun_imageE [elim!]: wenzelm@18372: assumes major: "c: (prod_fun f g)`r" wenzelm@18372: and cases: "!!x y. [| c=(f(x),g(y)); (x,y):r |] ==> P" wenzelm@18372: shows P wenzelm@18372: apply (rule major [THEN imageE]) wenzelm@18372: apply (rule_tac p = x in PairE) wenzelm@18372: apply (rule cases) wenzelm@18372: apply (blast intro: prod_fun) wenzelm@18372: apply blast wenzelm@18372: done wenzelm@11838: wenzelm@11838: oheimb@14101: constdefs oheimb@14101: upd_fst :: "('a => 'c) => 'a * 'b => 'c * 'b" oheimb@14101: "upd_fst f == prod_fun f id" oheimb@14101: oheimb@14101: upd_snd :: "('b => 'c) => 'a * 'b => 'a * 'c" oheimb@14101: "upd_snd f == prod_fun id f" oheimb@14101: oheimb@14101: lemma upd_fst_conv [simp]: "upd_fst f (x,y) = (f x,y)" wenzelm@18372: by (simp add: upd_fst_def) oheimb@14101: oheimb@14101: lemma upd_snd_conv [simp]: "upd_snd f (x,y) = (x,f y)" wenzelm@18372: by (simp add: upd_snd_def) oheimb@14101: wenzelm@11838: text {* wenzelm@11838: \bigskip Disjoint union of a family of sets -- Sigma. wenzelm@11838: *} wenzelm@11838: wenzelm@11838: lemma SigmaI [intro!]: "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" wenzelm@11838: by (unfold Sigma_def) blast wenzelm@11838: paulson@14952: lemma SigmaE [elim!]: wenzelm@11838: "[| c: Sigma A B; wenzelm@11838: !!x y.[| x:A; y:B(x); c=(x,y) |] ==> P wenzelm@11838: |] ==> P" wenzelm@11838: -- {* The general elimination rule. *} wenzelm@11838: by (unfold Sigma_def) blast wenzelm@11838: wenzelm@11838: text {* schirmer@15422: Elimination of @{term "(a, b) : A \ B"} -- introduces no wenzelm@11838: eigenvariables. wenzelm@11838: *} wenzelm@11838: wenzelm@11838: lemma SigmaD1: "(a, b) : Sigma A B ==> a : A" wenzelm@18372: by blast wenzelm@11838: wenzelm@11838: lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a" wenzelm@18372: by blast wenzelm@11838: wenzelm@11838: lemma SigmaE2: wenzelm@11838: "[| (a, b) : Sigma A B; wenzelm@11838: [| a:A; b:B(a) |] ==> P wenzelm@11838: |] ==> P" paulson@14952: by blast wenzelm@11838: paulson@14952: lemma Sigma_cong: schirmer@15422: "\A = B; !!x. x \ B \ C x = D x\ schirmer@15422: \ (SIGMA x: A. C x) = (SIGMA x: B. D x)" wenzelm@18372: by auto wenzelm@11838: wenzelm@11838: lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_empty1 [simp]: "Sigma {} B = {}" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_empty2 [simp]: "A <*> {} = {}" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV" wenzelm@11838: by auto wenzelm@11838: wenzelm@11838: lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)" wenzelm@11838: by auto wenzelm@11838: wenzelm@11838: lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV" wenzelm@11838: by auto wenzelm@11838: wenzelm@11838: lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)" wenzelm@11838: by (blast elim: equalityE) wenzelm@11838: wenzelm@11838: lemma SetCompr_Sigma_eq: wenzelm@11838: "Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: \bigskip Complex rules for Sigma. wenzelm@11838: *} wenzelm@11838: wenzelm@11838: lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma UN_Times_distrib: wenzelm@11838: "(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)" wenzelm@11838: -- {* Suggested by Pierre Chartier *} wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma split_paired_Ball_Sigma [simp]: wenzelm@11838: "(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma split_paired_Bex_Sigma [simp]: wenzelm@11838: "(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: text {* wenzelm@11838: Non-dependent versions are needed to avoid the need for higher-order wenzelm@11838: matching, especially when the rules are re-oriented. wenzelm@11838: *} wenzelm@11838: wenzelm@11838: lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)" wenzelm@11838: by blast wenzelm@11838: wenzelm@11838: oheimb@11493: lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A" wenzelm@11777: apply (rule_tac x = "(a, b)" in image_eqI) wenzelm@11777: apply auto wenzelm@11777: done wenzelm@11777: oheimb@11493: wenzelm@11838: text {* wenzelm@11838: Setup of internal @{text split_rule}. wenzelm@11838: *} wenzelm@11838: wenzelm@11032: constdefs wenzelm@11425: internal_split :: "('a => 'b => 'c) => 'a * 'b => 'c" wenzelm@11032: "internal_split == split" wenzelm@11032: wenzelm@11032: lemma internal_split_conv: "internal_split c (a, b) = c a b" wenzelm@11032: by (simp only: internal_split_def split_conv) wenzelm@11032: wenzelm@11032: hide const internal_split wenzelm@11032: oheimb@11025: use "Tools/split_rule.ML" wenzelm@11032: setup SplitRule.setup nipkow@10213: berghofe@15394: haftmann@21195: subsection {* Further lemmas *} haftmann@21195: haftmann@21195: lemma haftmann@21195: split_Pair: "split Pair x = x" haftmann@21195: unfolding split_def by auto haftmann@21195: haftmann@21195: lemma haftmann@21195: split_comp: "split (f \ g) x = f (g (fst x)) (snd x)" haftmann@21195: by (cases x, simp) haftmann@21195: haftmann@21195: berghofe@15394: subsection {* Code generator setup *} berghofe@15394: haftmann@21908: lemmas [code func] = fst_conv snd_conv haftmann@21908: haftmann@20588: instance unit :: eq .. haftmann@20588: haftmann@20588: lemma [code func]: haftmann@21454: "(u\unit) = v \ True" unfolding unit_eq [of u] unit_eq [of v] by rule+ haftmann@20588: haftmann@21908: code_type unit haftmann@21908: (SML "unit") haftmann@21908: (OCaml "unit") haftmann@21908: (Haskell "()") haftmann@21908: haftmann@20588: code_instance unit :: eq haftmann@20588: (Haskell -) haftmann@20588: haftmann@21908: code_const "op = \ unit \ unit \ bool" haftmann@21908: (Haskell infixl 4 "==") haftmann@21908: haftmann@21908: code_const Unity haftmann@21908: (SML "()") haftmann@21908: (OCaml "()") haftmann@21908: (Haskell "()") haftmann@21908: haftmann@21908: code_reserved SML haftmann@21908: unit haftmann@21908: haftmann@21908: code_reserved OCaml haftmann@21908: unit haftmann@21908: haftmann@20588: instance * :: (eq, eq) eq .. haftmann@20588: haftmann@20588: lemma [code func]: haftmann@21454: "(x1\'a\eq, y1\'b\eq) = (x2, y2) \ x1 = x2 \ y1 = y2" by auto haftmann@20588: haftmann@21908: code_type * haftmann@21908: (SML infix 2 "*") haftmann@21908: (OCaml infix 2 "*") haftmann@21908: (Haskell "!((_),/ (_))") haftmann@21908: haftmann@20588: code_instance * :: eq haftmann@20588: (Haskell -) haftmann@20588: haftmann@21908: code_const "op = \ 'a\eq \ 'b\eq \ 'a \ 'b \ bool" haftmann@20588: (Haskell infixl 4 "==") haftmann@20588: haftmann@21908: code_const Pair haftmann@21908: (SML "!((_),/ (_))") haftmann@21908: (OCaml "!((_),/ (_))") haftmann@21908: (Haskell "!((_),/ (_))") haftmann@20588: haftmann@22389: code_const fst and snd haftmann@22389: (Haskell "fst" and "snd") haftmann@22389: berghofe@15394: types_code berghofe@15394: "*" ("(_ */ _)") berghofe@16770: attach (term_of) {* berghofe@16770: fun term_of_id_42 f T g U (x, y) = HOLogic.pair_const T U $ f x $ g y; berghofe@16770: *} berghofe@16770: attach (test) {* berghofe@16770: fun gen_id_42 aG bG i = (aG i, bG i); berghofe@16770: *} berghofe@15394: berghofe@18706: consts_code berghofe@18706: "Pair" ("(_,/ _)") berghofe@18706: haftmann@21908: setup {* haftmann@21908: haftmann@21908: let haftmann@18013: haftmann@19039: fun strip_abs_split 0 t = ([], t) haftmann@19039: | strip_abs_split i (Abs (s, T, t)) = haftmann@18013: let haftmann@18013: val s' = Codegen.new_name t s; haftmann@18013: val v = Free (s', T) haftmann@19039: in apfst (cons v) (strip_abs_split (i-1) (subst_bound (v, t))) end haftmann@19039: | strip_abs_split i (u as Const ("split", _) $ t) = (case strip_abs_split (i+1) t of berghofe@15394: (v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u) berghofe@15394: | _ => ([], u)) haftmann@19039: | strip_abs_split i t = ([], t); haftmann@18013: berghofe@16634: fun let_codegen thy defs gr dep thyname brack t = (case strip_comb t of berghofe@16634: (t1 as Const ("Let", _), t2 :: t3 :: ts) => berghofe@15394: let berghofe@15394: fun dest_let (l as Const ("Let", _) $ t $ u) = haftmann@19039: (case strip_abs_split 1 u of berghofe@15394: ([p], u') => apfst (cons (p, t)) (dest_let u') berghofe@15394: | _ => ([], l)) berghofe@15394: | dest_let t = ([], t); berghofe@15394: fun mk_code (gr, (l, r)) = berghofe@15394: let berghofe@16634: val (gr1, pl) = Codegen.invoke_codegen thy defs dep thyname false (gr, l); berghofe@16634: val (gr2, pr) = Codegen.invoke_codegen thy defs dep thyname false (gr1, r); berghofe@15394: in (gr2, (pl, pr)) end berghofe@16634: in case dest_let (t1 $ t2 $ t3) of skalberg@15531: ([], _) => NONE berghofe@15394: | (ps, u) => berghofe@15394: let berghofe@15394: val (gr1, qs) = foldl_map mk_code (gr, ps); berghofe@16634: val (gr2, pu) = Codegen.invoke_codegen thy defs dep thyname false (gr1, u); berghofe@16634: val (gr3, pargs) = foldl_map berghofe@17021: (Codegen.invoke_codegen thy defs dep thyname true) (gr2, ts) berghofe@15394: in berghofe@16634: SOME (gr3, Codegen.mk_app brack berghofe@16634: (Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, List.concat berghofe@16634: (separate [Pretty.str ";", Pretty.brk 1] (map (fn (pl, pr) => berghofe@16634: [Pretty.block [Pretty.str "val ", pl, Pretty.str " =", berghofe@16634: Pretty.brk 1, pr]]) qs))), berghofe@16634: Pretty.brk 1, Pretty.str "in ", pu, berghofe@16634: Pretty.brk 1, Pretty.str "end"])) pargs) berghofe@15394: end berghofe@15394: end berghofe@16634: | _ => NONE); berghofe@15394: berghofe@16634: fun split_codegen thy defs gr dep thyname brack t = (case strip_comb t of berghofe@16634: (t1 as Const ("split", _), t2 :: ts) => haftmann@19039: (case strip_abs_split 1 (t1 $ t2) of berghofe@16634: ([p], u) => berghofe@16634: let berghofe@16634: val (gr1, q) = Codegen.invoke_codegen thy defs dep thyname false (gr, p); berghofe@16634: val (gr2, pu) = Codegen.invoke_codegen thy defs dep thyname false (gr1, u); berghofe@16634: val (gr3, pargs) = foldl_map berghofe@17021: (Codegen.invoke_codegen thy defs dep thyname true) (gr2, ts) berghofe@16634: in berghofe@16634: SOME (gr2, Codegen.mk_app brack berghofe@16634: (Pretty.block [Pretty.str "(fn ", q, Pretty.str " =>", berghofe@16634: Pretty.brk 1, pu, Pretty.str ")"]) pargs) berghofe@16634: end berghofe@16634: | _ => NONE) berghofe@16634: | _ => NONE); berghofe@15394: haftmann@21908: in haftmann@21908: haftmann@20105: Codegen.add_codegen "let_codegen" let_codegen haftmann@20105: #> Codegen.add_codegen "split_codegen" split_codegen haftmann@20105: #> CodegenPackage.add_appconst haftmann@20105: ("Let", CodegenPackage.appgen_let) berghofe@15394: haftmann@21908: end berghofe@15394: *} berghofe@15394: haftmann@21908: ML {* paulson@15404: val Collect_split = thm "Collect_split"; paulson@15404: val Compl_Times_UNIV1 = thm "Compl_Times_UNIV1"; paulson@15404: val Compl_Times_UNIV2 = thm "Compl_Times_UNIV2"; paulson@15404: val PairE = thm "PairE"; paulson@15404: val PairE_lemma = thm "PairE_lemma"; paulson@15404: val Pair_Rep_inject = thm "Pair_Rep_inject"; paulson@15404: val Pair_def = thm "Pair_def"; paulson@15404: val Pair_eq = thm "Pair_eq"; paulson@15404: val Pair_fst_snd_eq = thm "Pair_fst_snd_eq"; paulson@15404: val Pair_inject = thm "Pair_inject"; paulson@15404: val ProdI = thm "ProdI"; paulson@15404: val SetCompr_Sigma_eq = thm "SetCompr_Sigma_eq"; paulson@15404: val SigmaD1 = thm "SigmaD1"; paulson@15404: val SigmaD2 = thm "SigmaD2"; paulson@15404: val SigmaE = thm "SigmaE"; paulson@15404: val SigmaE2 = thm "SigmaE2"; paulson@15404: val SigmaI = thm "SigmaI"; paulson@15404: val Sigma_Diff_distrib1 = thm "Sigma_Diff_distrib1"; paulson@15404: val Sigma_Diff_distrib2 = thm "Sigma_Diff_distrib2"; paulson@15404: val Sigma_Int_distrib1 = thm "Sigma_Int_distrib1"; paulson@15404: val Sigma_Int_distrib2 = thm "Sigma_Int_distrib2"; paulson@15404: val Sigma_Un_distrib1 = thm "Sigma_Un_distrib1"; paulson@15404: val Sigma_Un_distrib2 = thm "Sigma_Un_distrib2"; paulson@15404: val Sigma_Union = thm "Sigma_Union"; paulson@15404: val Sigma_def = thm "Sigma_def"; paulson@15404: val Sigma_empty1 = thm "Sigma_empty1"; paulson@15404: val Sigma_empty2 = thm "Sigma_empty2"; paulson@15404: val Sigma_mono = thm "Sigma_mono"; paulson@15404: val The_split = thm "The_split"; paulson@15404: val The_split_eq = thm "The_split_eq"; paulson@15404: val The_split_eq = thm "The_split_eq"; paulson@15404: val Times_Diff_distrib1 = thm "Times_Diff_distrib1"; paulson@15404: val Times_Int_distrib1 = thm "Times_Int_distrib1"; paulson@15404: val Times_Un_distrib1 = thm "Times_Un_distrib1"; paulson@15404: val Times_eq_cancel2 = thm "Times_eq_cancel2"; paulson@15404: val Times_subset_cancel2 = thm "Times_subset_cancel2"; paulson@15404: val UNIV_Times_UNIV = thm "UNIV_Times_UNIV"; paulson@15404: val UN_Times_distrib = thm "UN_Times_distrib"; paulson@15404: val Unity_def = thm "Unity_def"; paulson@15404: val cond_split_eta = thm "cond_split_eta"; paulson@15404: val fst_conv = thm "fst_conv"; paulson@15404: val fst_def = thm "fst_def"; paulson@15404: val fst_eqD = thm "fst_eqD"; paulson@15404: val inj_on_Abs_Prod = thm "inj_on_Abs_Prod"; paulson@15404: val injective_fst_snd = thm "injective_fst_snd"; paulson@15404: val mem_Sigma_iff = thm "mem_Sigma_iff"; paulson@15404: val mem_splitE = thm "mem_splitE"; paulson@15404: val mem_splitI = thm "mem_splitI"; paulson@15404: val mem_splitI2 = thm "mem_splitI2"; paulson@15404: val prod_eqI = thm "prod_eqI"; paulson@15404: val prod_fun = thm "prod_fun"; paulson@15404: val prod_fun_compose = thm "prod_fun_compose"; paulson@15404: val prod_fun_def = thm "prod_fun_def"; paulson@15404: val prod_fun_ident = thm "prod_fun_ident"; paulson@15404: val prod_fun_imageE = thm "prod_fun_imageE"; paulson@15404: val prod_fun_imageI = thm "prod_fun_imageI"; paulson@15404: val prod_induct = thm "prod_induct"; paulson@15404: val snd_conv = thm "snd_conv"; paulson@15404: val snd_def = thm "snd_def"; paulson@15404: val snd_eqD = thm "snd_eqD"; paulson@15404: val split = thm "split"; paulson@15404: val splitD = thm "splitD"; paulson@15404: val splitD' = thm "splitD'"; paulson@15404: val splitE = thm "splitE"; paulson@15404: val splitE' = thm "splitE'"; paulson@15404: val splitE2 = thm "splitE2"; paulson@15404: val splitI = thm "splitI"; paulson@15404: val splitI2 = thm "splitI2"; paulson@15404: val splitI2' = thm "splitI2'"; paulson@15404: val split_Pair_apply = thm "split_Pair_apply"; paulson@15404: val split_beta = thm "split_beta"; paulson@15404: val split_conv = thm "split_conv"; paulson@15404: val split_def = thm "split_def"; paulson@15404: val split_eta = thm "split_eta"; paulson@15404: val split_eta_SetCompr = thm "split_eta_SetCompr"; paulson@15404: val split_eta_SetCompr2 = thm "split_eta_SetCompr2"; paulson@15404: val split_paired_All = thm "split_paired_All"; paulson@15404: val split_paired_Ball_Sigma = thm "split_paired_Ball_Sigma"; paulson@15404: val split_paired_Bex_Sigma = thm "split_paired_Bex_Sigma"; paulson@15404: val split_paired_Ex = thm "split_paired_Ex"; paulson@15404: val split_paired_The = thm "split_paired_The"; paulson@15404: val split_paired_all = thm "split_paired_all"; paulson@15404: val split_part = thm "split_part"; paulson@15404: val split_split = thm "split_split"; paulson@15404: val split_split_asm = thm "split_split_asm"; paulson@15404: val split_tupled_all = thms "split_tupled_all"; paulson@15404: val split_weak_cong = thm "split_weak_cong"; paulson@15404: val surj_pair = thm "surj_pair"; paulson@15404: val surjective_pairing = thm "surjective_pairing"; paulson@15404: val unit_abs_eta_conv = thm "unit_abs_eta_conv"; paulson@15404: val unit_all_eq1 = thm "unit_all_eq1"; paulson@15404: val unit_all_eq2 = thm "unit_all_eq2"; paulson@15404: val unit_eq = thm "unit_eq"; paulson@15404: val unit_induct = thm "unit_induct"; paulson@15404: *} paulson@15404: nipkow@10213: end