# HG changeset patch # User bulwahn # Date 1334221275 -7200 # Node ID d8fad618a67a3ac9a56de9f1619d638d390d710b # Parent ec64d94cbf9c91eaa9d623e326ef6dc9ea65fe13# Parent e1b761c216acc8a2c0663e426c0f91c2559f88ce merged diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Lifting.thy --- a/src/HOL/Lifting.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Lifting.thy Thu Apr 12 11:01:15 2012 +0200 @@ -199,19 +199,19 @@ apply safe apply (drule Abs1, simp) apply (erule Abs2) - apply (rule pred_compI) + apply (rule relcomppI) apply (rule Rep1) apply (rule Rep2) - apply (rule pred_compI, assumption) + apply (rule relcomppI, assumption) apply (drule Abs1, simp) apply (clarsimp simp add: R2) - apply (rule pred_compI, assumption) + apply (rule relcomppI, assumption) apply (drule Abs1, simp)+ apply (clarsimp simp add: R2) apply (drule Abs1, simp)+ apply (clarsimp simp add: R2) - apply (rule pred_compI, assumption) - apply (rule pred_compI [rotated]) + apply (rule relcomppI, assumption) + apply (rule relcomppI [rotated]) apply (erule conversepI) apply (drule Abs1, simp)+ apply (simp add: R2) diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/List.thy --- a/src/HOL/List.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/List.thy Thu Apr 12 11:01:15 2012 +0200 @@ -5677,7 +5677,7 @@ "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)" by (simp add: finite_trancl_ntranl) -lemma set_rel_comp [code]: +lemma set_relcomp [code]: "set xys O set yzs = set ([(fst xy, snd yz). xy \ xys, yz \ yzs, snd xy = fst yz])" by (auto simp add: Bex_def) diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy --- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Thu Apr 12 11:01:15 2012 +0200 @@ -1027,8 +1027,8 @@ (o * o => bool) => i * i => bool) [inductify] converse . thm converse.equation -code_pred [inductify] rel_comp . -thm rel_comp.equation +code_pred [inductify] relcomp . +thm relcomp.equation code_pred [inductify] Image . thm Image.equation declare singleton_iff[code_pred_inline] diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Quotient.thy --- a/src/HOL/Quotient.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Quotient.thy Thu Apr 12 11:01:15 2012 +0200 @@ -694,9 +694,9 @@ apply (rule Quotient3I) apply (simp add: o_def Quotient3_abs_rep [OF R2] Quotient3_abs_rep [OF R1]) apply simp - apply (rule_tac b="Rep1 (Rep2 a)" in pred_compI) + apply (rule_tac b="Rep1 (Rep2 a)" in relcomppI) apply (rule Quotient3_rep_reflp [OF R1]) - apply (rule_tac b="Rep1 (Rep2 a)" in pred_compI [rotated]) + apply (rule_tac b="Rep1 (Rep2 a)" in relcomppI [rotated]) apply (rule Quotient3_rep_reflp [OF R1]) apply (rule Rep1) apply (rule Quotient3_rep_reflp [OF R2]) @@ -707,8 +707,8 @@ apply (erule Quotient3_refl1 [OF R1]) apply (drule Quotient3_refl1 [OF R2], drule Rep1) apply (subgoal_tac "R1 r (Rep1 (Abs1 x))") - apply (rule_tac b="Rep1 (Abs1 x)" in pred_compI, assumption) - apply (erule pred_compI) + apply (rule_tac b="Rep1 (Abs1 x)" in relcomppI, assumption) + apply (erule relcomppI) apply (erule Quotient3_symp [OF R1, THEN sympD]) apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) apply (rule conjI, erule Quotient3_refl1 [OF R1]) @@ -721,8 +721,8 @@ apply (erule Quotient3_refl1 [OF R1]) apply (drule Quotient3_refl2 [OF R2], drule Rep1) apply (subgoal_tac "R1 s (Rep1 (Abs1 y))") - apply (rule_tac b="Rep1 (Abs1 y)" in pred_compI, assumption) - apply (erule pred_compI) + apply (rule_tac b="Rep1 (Abs1 y)" in relcomppI, assumption) + apply (erule relcomppI) apply (erule Quotient3_symp [OF R1, THEN sympD]) apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) apply (rule conjI, erule Quotient3_refl2 [OF R1]) @@ -738,11 +738,11 @@ apply (erule Quotient3_refl1 [OF R1]) apply (rename_tac a b c d) apply simp - apply (rule_tac b="Rep1 (Abs1 r)" in pred_compI) + apply (rule_tac b="Rep1 (Abs1 r)" in relcomppI) apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) apply (rule conjI, erule Quotient3_refl1 [OF R1]) apply (simp add: Quotient3_abs_rep [OF R1] Quotient3_rep_reflp [OF R1]) - apply (rule_tac b="Rep1 (Abs1 s)" in pred_compI [rotated]) + apply (rule_tac b="Rep1 (Abs1 s)" in relcomppI [rotated]) apply (rule Quotient3_rel[symmetric, OF R1, THEN iffD2]) apply (simp add: Quotient3_abs_rep [OF R1] Quotient3_rep_reflp [OF R1]) apply (erule Quotient3_refl2 [OF R1]) diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Quotient_Examples/FSet.thy --- a/src/HOL/Quotient_Examples/FSet.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Quotient_Examples/FSet.thy Thu Apr 12 11:01:15 2012 +0200 @@ -140,7 +140,7 @@ next assume a: "(list_all2 R OOO op \) r s" then have b: "map Abs r \ map Abs s" - proof (elim pred_compE) + proof (elim relcomppE) fix b ba assume c: "list_all2 R r b" assume d: "b \ ba" @@ -165,11 +165,11 @@ have y: "list_all2 R (map Rep (map Abs s)) s" by (fact rep_abs_rsp_left[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of s]]) have c: "(op \ OO list_all2 R) (map Rep (map Abs r)) s" - by (rule pred_compI) (rule b, rule y) + by (rule relcomppI) (rule b, rule y) have z: "list_all2 R r (map Rep (map Abs r))" by (fact rep_abs_rsp[OF list_quotient3[OF q], OF list_all2_refl'[OF e, of r]]) then show "(list_all2 R OOO op \) r s" - using a c pred_compI by simp + using a c relcomppI by simp qed qed @@ -360,7 +360,7 @@ quotient_definition "concat_fset :: ('a fset) fset \ 'a fset" is concat -proof (elim pred_compE) +proof (elim relcomppE) fix a b ba bb assume a: "list_all2 op \ a ba" with list_symp [OF list_eq_symp] have a': "list_all2 op \ ba a" by (rule sympE) @@ -397,9 +397,9 @@ lemma Cons_rsp2 [quot_respect]: shows "(op \ ===> list_all2 op \ OOO op \ ===> list_all2 op \ OOO op \) Cons Cons" apply (auto intro!: fun_relI) - apply (rule_tac b="x # b" in pred_compI) + apply (rule_tac b="x # b" in relcomppI) apply auto - apply (rule_tac b="x # ba" in pred_compI) + apply (rule_tac b="x # ba" in relcomppI) apply auto done @@ -453,7 +453,7 @@ assumes "(R1 ===> R2 ===> R3) C C" and "(R4 ===> R5 ===> R6) C C" shows "(R1 OOO R4 ===> R2 OOO R5 ===> R3 OOO R6) C C" by (auto intro!: fun_relI) - (metis (full_types) assms fun_relE pred_compI) + (metis (full_types) assms fun_relE relcomppI) lemma append_rsp2 [quot_respect]: "(list_all2 op \ OOO op \ ===> list_all2 op \ OOO op \ ===> list_all2 op \ OOO op \) append append" @@ -479,7 +479,7 @@ by (induct y ya rule: list_induct2') (simp_all, metis apply_rsp' list_eq_def) show "(list_all2 op \ OOO op \) (map f xa) (map f' ya)" - by (metis a b c list_eq_def pred_compI) + by (metis a b c list_eq_def relcomppI) qed lemma map_prs2 [quot_preserve]: diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Relation.thy --- a/src/HOL/Relation.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Relation.thy Thu Apr 12 11:01:15 2012 +0200 @@ -507,29 +507,26 @@ subsubsection {* Composition *} -inductive_set rel_comp :: "('a \ 'b) set \ ('b \ 'c) set \ ('a \ 'c) set" (infixr "O" 75) +inductive_set relcomp :: "('a \ 'b) set \ ('b \ 'c) set \ ('a \ 'c) set" (infixr "O" 75) for r :: "('a \ 'b) set" and s :: "('b \ 'c) set" where - rel_compI [intro]: "(a, b) \ r \ (b, c) \ s \ (a, c) \ r O s" + relcompI [intro]: "(a, b) \ r \ (b, c) \ s \ (a, c) \ r O s" -abbreviation pred_comp (infixr "OO" 75) where - "pred_comp \ rel_compp" +notation relcompp (infixr "OO" 75) -lemmas pred_compI = rel_compp.intros +lemmas relcomppI = relcompp.intros text {* For historic reasons, the elimination rules are not wholly corresponding. Feel free to consolidate this. *} -inductive_cases rel_compEpair: "(a, c) \ r O s" -inductive_cases pred_compE [elim!]: "(r OO s) a c" +inductive_cases relcompEpair: "(a, c) \ r O s" +inductive_cases relcomppE [elim!]: "(r OO s) a c" -lemma rel_compE [elim!]: "xz \ r O s \ +lemma relcompE [elim!]: "xz \ r O s \ (\x y z. xz = (x, z) \ (x, y) \ r \ (y, z) \ s \ P) \ P" - by (cases xz) (simp, erule rel_compEpair, iprover) - -lemmas pred_comp_rel_comp_eq = rel_compp_rel_comp_eq + by (cases xz) (simp, erule relcompEpair, iprover) lemma R_O_Id [simp]: "R O Id = R" @@ -539,28 +536,28 @@ "Id O R = R" by fast -lemma rel_comp_empty1 [simp]: +lemma relcomp_empty1 [simp]: "{} O R = {}" by blast -lemma pred_comp_bot1 [simp]: +lemma relcompp_bot1 [simp]: "\ OO R = \" - by (fact rel_comp_empty1 [to_pred]) + by (fact relcomp_empty1 [to_pred]) -lemma rel_comp_empty2 [simp]: +lemma relcomp_empty2 [simp]: "R O {} = {}" by blast -lemma pred_comp_bot2 [simp]: +lemma relcompp_bot2 [simp]: "R OO \ = \" - by (fact rel_comp_empty2 [to_pred]) + by (fact relcomp_empty2 [to_pred]) lemma O_assoc: "(R O S) O T = R O (S O T)" by blast -lemma pred_comp_assoc: +lemma relcompp_assoc: "(r OO s) OO t = r OO (s OO t)" by (fact O_assoc [to_pred]) @@ -568,55 +565,55 @@ "trans r \ r O r \ r" by (unfold trans_def) blast -lemma transp_pred_comp_less_eq: +lemma transp_relcompp_less_eq: "transp r \ r OO r \ r " by (fact trans_O_subset [to_pred]) -lemma rel_comp_mono: +lemma relcomp_mono: "r' \ r \ s' \ s \ r' O s' \ r O s" by blast -lemma pred_comp_mono: +lemma relcompp_mono: "r' \ r \ s' \ s \ r' OO s' \ r OO s " - by (fact rel_comp_mono [to_pred]) + by (fact relcomp_mono [to_pred]) -lemma rel_comp_subset_Sigma: +lemma relcomp_subset_Sigma: "r \ A \ B \ s \ B \ C \ r O s \ A \ C" by blast -lemma rel_comp_distrib [simp]: +lemma relcomp_distrib [simp]: "R O (S \ T) = (R O S) \ (R O T)" by auto -lemma pred_comp_distrib [simp]: +lemma relcompp_distrib [simp]: "R OO (S \ T) = R OO S \ R OO T" - by (fact rel_comp_distrib [to_pred]) + by (fact relcomp_distrib [to_pred]) -lemma rel_comp_distrib2 [simp]: +lemma relcomp_distrib2 [simp]: "(S \ T) O R = (S O R) \ (T O R)" by auto -lemma pred_comp_distrib2 [simp]: +lemma relcompp_distrib2 [simp]: "(S \ T) OO R = S OO R \ T OO R" - by (fact rel_comp_distrib2 [to_pred]) + by (fact relcomp_distrib2 [to_pred]) -lemma rel_comp_UNION_distrib: +lemma relcomp_UNION_distrib: "s O UNION I r = (\i\I. s O r i) " by auto -(* FIXME thm rel_comp_UNION_distrib [to_pred] *) +(* FIXME thm relcomp_UNION_distrib [to_pred] *) -lemma rel_comp_UNION_distrib2: +lemma relcomp_UNION_distrib2: "UNION I r O s = (\i\I. r i O s) " by auto -(* FIXME thm rel_comp_UNION_distrib2 [to_pred] *) +(* FIXME thm relcomp_UNION_distrib2 [to_pred] *) -lemma single_valued_rel_comp: +lemma single_valued_relcomp: "single_valued r \ single_valued s \ single_valued (r O s)" by (unfold single_valued_def) blast -lemma rel_comp_unfold: +lemma relcomp_unfold: "r O s = {(x, z). \y. (x, y) \ r \ (y, z) \ s}" by (auto simp add: set_eq_iff) @@ -676,12 +673,12 @@ "(r\\)\\ = r" by (fact converse_converse [to_pred]) -lemma converse_rel_comp: "(r O s)^-1 = s^-1 O r^-1" +lemma converse_relcomp: "(r O s)^-1 = s^-1 O r^-1" by blast -lemma converse_pred_comp: "(r OO s)^--1 = s^--1 OO r^--1" - by (iprover intro: order_antisym conversepI pred_compI - elim: pred_compE dest: conversepD) +lemma converse_relcompp: "(r OO s)^--1 = s^--1 OO r^--1" + by (iprover intro: order_antisym conversepI relcomppI + elim: relcomppE dest: conversepD) lemma converse_Int: "(r \ s)^-1 = r^-1 \ s^-1" by blast diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Tools/Function/termination.ML --- a/src/HOL/Tools/Function/termination.ML Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Tools/Function/termination.ML Thu Apr 12 11:01:15 2012 +0200 @@ -141,7 +141,7 @@ fun prove_chain thy chain_tac (c1, c2) = let val goal = - HOLogic.mk_eq (HOLogic.mk_binop @{const_name Relation.rel_comp} (c1, c2), + HOLogic.mk_eq (HOLogic.mk_binop @{const_name Relation.relcomp} (c1, c2), Const (@{const_abbrev Set.empty}, fastype_of c1)) |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *) in diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Thu Apr 12 11:01:15 2012 +0200 @@ -391,7 +391,7 @@ (@{const_name Id}, 0), (@{const_name converse}, 1), (@{const_name trancl}, 1), - (@{const_name rel_comp}, 2), + (@{const_name relcomp}, 2), (@{const_name finite}, 1), (@{const_name unknown}, 0), (@{const_name is_unknown}, 1), diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Tools/Nitpick/nitpick_mono.ML --- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Thu Apr 12 11:01:15 2012 +0200 @@ -856,7 +856,7 @@ accum ||> add_annotation_atom_comp Neq [] (V x) (A New)) end | @{const_name trancl} => do_fragile_set_operation T accum - | @{const_name rel_comp} => + | @{const_name relcomp} => let val x = Unsynchronized.inc max_fresh val bc_set_M = domain_type T |> mtype_for_set x diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Tools/Nitpick/nitpick_nut.ML --- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Thu Apr 12 11:01:15 2012 +0200 @@ -522,7 +522,7 @@ Op1 (Converse, range_type T, Any, sub t1) | (Const (@{const_name trancl}, T), [t1]) => Op1 (Closure, range_type T, Any, sub t1) - | (Const (@{const_name rel_comp}, T), [t1, t2]) => + | (Const (@{const_name relcomp}, T), [t1, t2]) => Op2 (Composition, nth_range_type 2 T, Any, sub t1, sub t2) | (Const (x as (s as @{const_name Suc}, T)), []) => if is_built_in_const thy stds x then Cst (Suc, T, Any) diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Transitive_Closure.thy Thu Apr 12 11:01:15 2012 +0200 @@ -726,7 +726,7 @@ lemma relpowp_relpow_eq [pred_set_conv]: fixes R :: "'a rel" shows "(\x y. (x, y) \ R) ^^ n = (\x y. (x, y) \ R ^^ n)" - by (induct n) (simp_all add: rel_compp_rel_comp_eq) + by (induct n) (simp_all add: relcompp_relcomp_eq) text {* for code generation *} @@ -849,7 +849,7 @@ apply (drule tranclD2) apply (clarsimp simp: rtrancl_is_UN_relpow) apply (rule_tac x="Suc n" in exI) - apply (clarsimp simp: rel_comp_unfold) + apply (clarsimp simp: relcomp_unfold) apply fastforce apply clarsimp apply (case_tac n, simp) @@ -870,7 +870,7 @@ next case (Suc n) show ?case - proof (simp add: rel_comp_unfold Suc) + proof (simp add: relcomp_unfold Suc) show "(\y. (\f. f 0 = a \ f n = y \ (\i R)) \ (y,b) \ R) = (\f. f 0 = a \ f(Suc n) = b \ (\i R))" (is "?l = ?r") @@ -979,7 +979,7 @@ apply(auto dest: relpow_finite_bounded1) done -lemma finite_rel_comp[simp,intro]: +lemma finite_relcomp[simp,intro]: assumes "finite R" and "finite S" shows "finite(R O S)" proof- diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/Wellfounded.thy Thu Apr 12 11:01:15 2012 +0200 @@ -271,7 +271,7 @@ assume "y \ Q" with `y \ ?Q'` obtain w where "(w, y) \ R" and "w \ Q" by auto - from `(w, y) \ R` `(y, z) \ S` have "(w, z) \ R O S" by (rule rel_compI) + from `(w, y) \ R` `(y, z) \ S` have "(w, z) \ R O S" by (rule relcompI) with `R O S \ R` have "(w, z) \ R" .. with `z \ ?Q'` have "w \ Q" by blast with `w \ Q` show False by contradiction diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/ex/Coherent.thy --- a/src/HOL/ex/Coherent.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/ex/Coherent.thy Thu Apr 12 11:01:15 2012 +0200 @@ -13,7 +13,7 @@ no_notation comp (infixl "o" 55) and - rel_comp (infixr "O" 75) + relcomp (infixr "O" 75) lemma p1p2: assumes diff -r ec64d94cbf9c -r d8fad618a67a src/HOL/ex/Executable_Relation.thy --- a/src/HOL/ex/Executable_Relation.thy Thu Apr 12 10:13:33 2012 +0200 +++ b/src/HOL/ex/Executable_Relation.thy Thu Apr 12 11:01:15 2012 +0200 @@ -27,7 +27,7 @@ lemma comp_Id_on: "Id_on X O R = Set.project (%(x, y). x : X) R" -by (auto intro!: rel_compI) +by (auto intro!: relcompI) lemma comp_Id_on': "R O Id_on X = Set.project (%(x, y). y : X) R" @@ -37,7 +37,7 @@ "Set.project (%(x, y). x : X) (Id_on Y) = Id_on (X Int Y)" by auto -lemma rel_comp_raw: +lemma relcomp_raw: "(rel_raw X R) O (rel_raw Y S) = rel_raw (X Int Y) (Set.project (%(x, y). y : Y) R Un (Set.project (%(x, y). x : X) S Un R O S))" unfolding rel_raw_def apply simp @@ -79,7 +79,7 @@ subsubsection {* Constant definitions on relations *} -hide_const (open) converse rel_comp rtrancl Image +hide_const (open) converse relcomp rtrancl Image quotient_definition member :: "'a * 'a => 'a rel => bool" where "member" is "Set.member :: 'a * 'a => ('a * 'a) set => bool" done @@ -92,9 +92,9 @@ where "union" is "Set.union :: ('a * 'a) set => ('a * 'a) set => ('a * 'a) set" done -quotient_definition rel_comp :: "'a rel => 'a rel => 'a rel" +quotient_definition relcomp :: "'a rel => 'a rel => 'a rel" where - "rel_comp" is "Relation.rel_comp :: ('a * 'a) set => ('a * 'a) set => ('a * 'a) set" done + "relcomp" is "Relation.relcomp :: ('a * 'a) set => ('a * 'a) set => ('a * 'a) set" done quotient_definition rtrancl :: "'a rel => 'a rel" where @@ -121,8 +121,8 @@ by (lifting union_raw) lemma [code]: - "rel_comp (rel X R) (rel Y S) = rel (X Int Y) (Set.project (%(x, y). y : Y) R Un (Set.project (%(x, y). x : X) S Un R O S))" -by (lifting rel_comp_raw) + "relcomp (rel X R) (rel Y S) = rel (X Int Y) (Set.project (%(x, y). y : Y) R Un (Set.project (%(x, y). x : X) S Un R O S))" +by (lifting relcomp_raw) lemma [code]: "rtrancl (rel X R) = rel UNIV (R^+)"