# HG changeset patch # User haftmann # Date 1240212760 -7200 # Node ID cf50e67bc1d1b351a8ea2b21c66b9b7057095fb6 # Parent d5f5ab29d76907e5fa9a042380e9edf3b702322b power operation on functions in theory Nat; power operation on relations in theory Transitive_Closure diff -r d5f5ab29d769 -r cf50e67bc1d1 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Apr 20 09:32:09 2009 +0200 +++ b/src/HOL/IsaMakefile Mon Apr 20 09:32:40 2009 +0200 @@ -218,7 +218,6 @@ Nat_Numeral.thy \ Presburger.thy \ Recdef.thy \ - Relation_Power.thy \ SetInterval.thy \ $(SRC)/Provers/Arith/assoc_fold.ML \ $(SRC)/Provers/Arith/cancel_numeral_factor.ML \ diff -r d5f5ab29d769 -r cf50e67bc1d1 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Mon Apr 20 09:32:09 2009 +0200 +++ b/src/HOL/Nat.thy Mon Apr 20 09:32:40 2009 +0200 @@ -1164,6 +1164,37 @@ end +subsection {* Natural operation of natural numbers on functions *} + +text {* @{text "f o^ n = f o ... o f"}, the n-fold composition of @{text f} *} + +primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where + "funpow 0 f = id" + | "funpow (Suc n) f = f o funpow n f" + +abbreviation funpower :: "('a \ 'a) \ nat \ 'a \ 'a" (infixr "o^" 80) where + "f o^ n \ funpow n f" + +notation (latex output) + funpower ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +notation (HTML output) + funpower ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +lemma funpow_add: + "f o^ (m + n) = f o^ m \ f o^ n" + by (induct m) simp_all + +lemma funpow_swap1: + "f ((f o^ n) x) = (f o^ n) (f x)" +proof - + have "f ((f o^ n) x) = (f o^ (n + 1)) x" by simp + also have "\ = (f o^ n o f o^ 1) x" by (simp only: funpow_add) + also have "\ = (f o^ n) (f x)" by simp + finally show ?thesis . +qed + + subsection {* Embedding of the Naturals into any @{text semiring_1}: @{term of_nat} *} diff -r d5f5ab29d769 -r cf50e67bc1d1 src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Mon Apr 20 09:32:09 2009 +0200 +++ b/src/HOL/Transitive_Closure.thy Mon Apr 20 09:32:40 2009 +0200 @@ -630,6 +630,139 @@ declare trancl_into_rtrancl [elim] +subsection {* The power operation on relations *} + +text {* @{text "R ^^ n = R O ... O R"}, the n-fold composition of @{text R} *} + +primrec relpow :: "('a \ 'a) set \ nat \ ('a \ 'a) set" (infixr "^^" 80) where + "R ^^ 0 = Id" + | "R ^^ Suc n = R O (R ^^ n)" + +notation (latex output) + relpow ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +notation (HTML output) + relpow ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +lemma rel_pow_1 [simp]: + "R ^^ 1 = R" + by simp + +lemma rel_pow_0_I: + "(x, x) \ R ^^ 0" + by simp + +lemma rel_pow_Suc_I: + "(x, y) \ R ^^ n \ (y, z) \ R \ (x, z) \ R ^^ Suc n" + by auto + +lemma rel_pow_Suc_I2: + "(x, y) \ R \ (y, z) \ R ^^ n \ (x, z) \ R ^^ Suc n" + by (induct n arbitrary: z) (simp, fastsimp) + +lemma rel_pow_0_E: + "(x, y) \ R ^^ 0 \ (x = y \ P) \ P" + by simp + +lemma rel_pow_Suc_E: + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R ^^ n \ (y, z) \ R \ P) \ P" + by auto + +lemma rel_pow_E: + "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) + \ (\y m. n = Suc m \ (x, y) \ R ^^ m \ (y, z) \ R \ P) + \ P" + by (cases n) auto + +lemma rel_pow_Suc_D2: + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n)" + apply (induct n arbitrary: x z) + apply (blast intro: rel_pow_0_I elim: rel_pow_0_E rel_pow_Suc_E) + apply (blast intro: rel_pow_Suc_I elim: rel_pow_0_E rel_pow_Suc_E) + done + +lemma rel_pow_Suc_E2: + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n \ P) \ P" + by (blast dest: rel_pow_Suc_D2) + +lemma rel_pow_Suc_D2': + "\x y z. (x, y) \ R ^^ n \ (y, z) \ R \ (\w. (x, w) \ R \ (w, z) \ R ^^ n)" + by (induct n) (simp_all, blast) + +lemma rel_pow_E2: + "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) + \ (\y m. n = Suc m \ (x, y) \ R \ (y, z) \ R ^^ m \ P) + \ P" + apply (cases n, simp) + apply (cut_tac n=nat and R=R in rel_pow_Suc_D2', simp, blast) + done + +lemma rtrancl_imp_UN_rel_pow: + assumes "p \ R^*" + shows "p \ (\n. R ^^ n)" +proof (cases p) + case (Pair x y) + with assms have "(x, y) \ R^*" by simp + then have "(x, y) \ (\n. R ^^ n)" proof induct + case base show ?case by (blast intro: rel_pow_0_I) + next + case step then show ?case by (blast intro: rel_pow_Suc_I) + qed + with Pair show ?thesis by simp +qed + +lemma rel_pow_imp_rtrancl: + assumes "p \ R ^^ n" + shows "p \ R^*" +proof (cases p) + case (Pair x y) + with assms have "(x, y) \ R ^^ n" by simp + then have "(x, y) \ R^*" proof (induct n arbitrary: x y) + case 0 then show ?case by simp + next + case Suc then show ?case + by (blast elim: rel_pow_Suc_E intro: rtrancl_into_rtrancl) + qed + with Pair show ?thesis by simp +qed + +lemma rtrancl_is_UN_rel_pow: + "R^* = (\n. R ^^ n)" + by (blast intro: rtrancl_imp_UN_rel_pow rel_pow_imp_rtrancl) + +lemma rtrancl_power: + "p \ R^* \ (\n. p \ R ^^ n)" + by (simp add: rtrancl_is_UN_rel_pow) + +lemma trancl_power: + "p \ R^+ \ (\n > 0. p \ R ^^ n)" + apply (cases p) + apply simp + apply (rule iffI) + apply (drule tranclD2) + apply (clarsimp simp: rtrancl_is_UN_rel_pow) + apply (rule_tac x="Suc x" in exI) + apply (clarsimp simp: rel_comp_def) + apply fastsimp + apply clarsimp + apply (case_tac n, simp) + apply clarsimp + apply (drule rel_pow_imp_rtrancl) + apply (drule rtrancl_into_trancl1) apply auto + done + +lemma rtrancl_imp_rel_pow: + "p \ R^* \ \n. p \ R ^^ n" + by (auto dest: rtrancl_imp_UN_rel_pow) + +lemma single_valued_rel_pow: + fixes R :: "('a * 'a) set" + shows "single_valued R \ single_valued (R ^^ n)" + apply (induct n arbitrary: R) + apply simp_all + apply (rule single_valuedI) + apply (fast dest: single_valuedD elim: rel_pow_Suc_E) + done subsection {* Setup of transitivity reasoner *}