nipkow@10213: (* Title: HOL/Transitive_Closure.thy nipkow@10213: ID: $Id$ nipkow@10213: Author: Lawrence C Paulson, Cambridge University Computer Laboratory nipkow@10213: Copyright 1992 University of Cambridge nipkow@10213: nipkow@10213: Relfexive and Transitive closure of a relation nipkow@10213: nipkow@10213: rtrancl is reflexive/transitive closure; nipkow@10213: trancl is transitive closure nipkow@10213: reflcl is reflexive closure nipkow@10213: wenzelm@10331: These postfix operators have MAXIMUM PRIORITY, forcing their operands wenzelm@10331: to be atomic. nipkow@10213: *) nipkow@10213: wenzelm@10980: theory Transitive_Closure = Lfp + Relation wenzelm@10980: files ("Transitive_Closure_lemmas.ML"): nipkow@10213: nipkow@10213: constdefs wenzelm@10331: rtrancl :: "('a * 'a) set => ('a * 'a) set" ("(_^*)" [1000] 999) wenzelm@10331: "r^* == lfp (%s. Id Un (r O s))" nipkow@10213: wenzelm@10331: trancl :: "('a * 'a) set => ('a * 'a) set" ("(_^+)" [1000] 999) wenzelm@10331: "r^+ == r O rtrancl r" nipkow@10213: nipkow@10213: syntax wenzelm@10331: "_reflcl" :: "('a * 'a) set => ('a * 'a) set" ("(_^=)" [1000] 999) nipkow@10213: translations nipkow@10213: "r^=" == "r Un Id" nipkow@10213: wenzelm@10827: syntax (xsymbols) wenzelm@10331: rtrancl :: "('a * 'a) set => ('a * 'a) set" ("(_\\<^sup>*)" [1000] 999) wenzelm@10331: trancl :: "('a * 'a) set => ('a * 'a) set" ("(_\\<^sup>+)" [1000] 999) wenzelm@10331: "_reflcl" :: "('a * 'a) set => ('a * 'a) set" ("(_\\<^sup>=)" [1000] 999) wenzelm@10331: wenzelm@10980: use "Transitive_Closure_lemmas.ML" wenzelm@10980: nipkow@10996: nipkow@10996: lemma reflcl_trancl[simp]: "(r\<^sup>+)\<^sup>= = r\<^sup>*" nipkow@10996: apply safe; nipkow@10996: apply (erule trancl_into_rtrancl); nipkow@10996: by (blast elim:rtranclE dest:rtrancl_into_trancl1) nipkow@10996: nipkow@10996: lemma trancl_reflcl[simp]: "(r\<^sup>=)\<^sup>+ = r\<^sup>*" nipkow@10996: apply safe nipkow@10996: apply (drule trancl_into_rtrancl) nipkow@10996: apply simp; nipkow@10996: apply (erule rtranclE) nipkow@10996: apply safe nipkow@10996: apply(rule r_into_trancl) nipkow@10996: apply simp nipkow@10996: apply(rule rtrancl_into_trancl1) nipkow@10996: apply(erule rtrancl_reflcl[THEN equalityD2, THEN subsetD]) nipkow@10996: apply fast nipkow@10996: done nipkow@10996: nipkow@10996: lemma trancl_empty[simp]: "{}\<^sup>+ = {}" nipkow@10996: by (auto elim:trancl_induct) nipkow@10996: nipkow@10996: lemma rtrancl_empty[simp]: "{}\<^sup>* = Id" nipkow@10996: by(rule subst[OF reflcl_trancl], simp) nipkow@10996: nipkow@10996: lemma rtranclD: "(a,b) \ R\<^sup>* \ a=b \ a\b \ (a,b) \ R\<^sup>+" nipkow@10996: by(force simp add: reflcl_trancl[THEN sym] simp del: reflcl_trancl) nipkow@10996: nipkow@10996: (* should be merged with the main body of lemmas: *) nipkow@10996: nipkow@10996: lemma Domain_rtrancl[simp]: "Domain(R\<^sup>*) = UNIV" nipkow@10996: by blast nipkow@10996: nipkow@10996: lemma Range_rtrancl[simp]: "Range(R\<^sup>*) = UNIV" nipkow@10996: by blast nipkow@10996: nipkow@10996: lemma rtrancl_Un_subset: "(R\<^sup>* \ S\<^sup>*) \ (R Un S)\<^sup>*" nipkow@10996: by(rule rtrancl_Un_rtrancl[THEN subst], fast) nipkow@10996: nipkow@10996: lemma in_rtrancl_UnI: "x \ R\<^sup>* \ x \ S\<^sup>* \ x \ (R \ S)\<^sup>*" nipkow@10996: by (blast intro: subsetD[OF rtrancl_Un_subset]) nipkow@10996: nipkow@10996: lemma trancl_domain [simp]: "Domain (r\<^sup>+) = Domain r" nipkow@10996: by (unfold Domain_def, blast dest:tranclD) nipkow@10996: nipkow@10996: lemma trancl_range [simp]: "Range (r\<^sup>+) = Range r" nipkow@10996: by (simp add:Range_def trancl_converse[THEN sym]) nipkow@10996: nipkow@10213: end