Trancl.thy
author lcp
Wed, 22 Sep 1993 15:43:05 +0200
changeset 2 befa4e9f7c90
parent 0 7949f97df77a
child 120 19facfd773de
permissions -rw-r--r--
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong, split_weak_cong, nat_case_weak_cong, nat_rec_weak_cong. Used in llist.ML to make simplifications faster. HOL/gfp: re-ordered premises to put mono(f) early (first or right after A==gfp(f) in the def_ rules). Renamed some variables in rules, A to X and h to A. Renamed coinduct to weak_coinduct, coinduct2 to coinduct. Strengthened coinduct as suggested by j. Frost, to have the premise X <= f(X Un gfp(f)). HOL/llist: used stronger coinduct rule to strengthen LList_coinduct, LList_equalityI, llist_equalityI, llist_fun_equalityI and to delete the "2" form of those rules. Proved List_Fun_LList_I, LListD_Fun_diag_I and llistD_Fun_range_I to help use the new coinduction rules; most proofs involving them required small changes. Proved prod_fun_range_eq_diag as lemma for llist_equalityI.

(*  Title: 	HOL/trancl.thy
    ID:         $Id$
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1992  University of Cambridge

Transitive closure of a relation

rtrancl is refl/transitive closure;  trancl is transitive closure
*)

Trancl = Lfp +
consts
    trans   :: "('a * 'a)set => bool" 	(*transitivity predicate*)
    id	    :: "('a * 'a)set"
    rtrancl :: "('a * 'a)set => ('a * 'a)set"	("(_^*)" [100] 100)
    trancl  :: "('a * 'a)set => ('a * 'a)set"	("(_^+)" [100] 100)  
    O	    :: "[('b * 'c)set, ('a * 'b)set] => ('a * 'c)set" (infixr 60)
rules   
trans_def	"trans(r) == (!x y z. <x,y>:r --> <y,z>:r --> <x,z>:r)"
comp_def	(*composition of relations*)
		"r O s == {xz. ? x y z. xz = <x,z> & <x,y>:s & <y,z>:r}"
id_def		(*the identity relation*)
		"id == {p. ? x. p = <x,x>}"
rtrancl_def	"r^* == lfp(%s. id Un (r O s))"
trancl_def	"r^+ == r O rtrancl(r)"
end