# HG changeset patch # User wenzelm # Date 971904482 -7200 # Node ID d99e5eeb16f4c05b7ef70d41be3d8b530f156ae2 # Parent 33e542b4934c2a02a1dca8c8fc6f952a5be9a970 The accessible part of a relation (from HOL/Induct/Acc); diff -r 33e542b4934c -r d99e5eeb16f4 src/HOL/Library/Accessible_Part.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Accessible_Part.ML Wed Oct 18 23:28:02 2000 +0200 @@ -0,0 +1,8 @@ + +val accI = thm "acc.accI"; +val acc_induct = thm "acc_induct"; +val acc_downward = thm "acc_downward"; +val acc_downwards = thm "acc_downwards"; +val acc_wfI = thm "acc_wfI"; +val acc_wfD = thm "acc_wfD"; +val wf_acc_iff = thm "wf_acc_iff"; diff -r 33e542b4934c -r d99e5eeb16f4 src/HOL/Library/Accessible_Part.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Accessible_Part.thy Wed Oct 18 23:28:02 2000 +0200 @@ -0,0 +1,83 @@ +(* Title: HOL/Library/Accessible_Part.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge +*) + +header {* + \title{The accessible part of a relation} + \author{Lawrence C Paulson} +*} + +theory Accessible_Part = Main: + + +subsection {* Inductive definition *} + +text {* + Inductive definition of the accessible part @{term "acc r"} of a + relation; see also \cite{paulin-tlca}. +*} + +consts + acc :: "('a \ 'a) set => 'a set" +inductive "acc r" + intros + accI [rule_format]: "\y. (y, x) \ r --> y \ acc r ==> x \ acc r" + +syntax + termi :: "('a \ 'a) set => 'a set" +translations + "termi r" == "acc (r^-1)" + + +subsection {* Induction rules *} + +theorem acc_induct [induct set: acc]: + "a \ acc r ==> + (!!x. x \ acc r ==> \y. (y, x) \ r --> P y ==> P x) ==> P a" +proof - + assume major: "a \ acc r" + assume hyp: "!!x. x \ acc r ==> \y. (y, x) \ r --> P y ==> P x" + show ?thesis + apply (rule major [THEN acc.induct]) + apply (rule hyp) + apply (rule accI) + apply fast + apply fast + done +qed + +theorem acc_downward: "b \ acc r ==> (a, b) \ r ==> a \ acc r" + apply (erule acc.elims) + apply fast + done + +lemma acc_downwards_aux: "(b, a) \ r^* ==> a \ acc r --> b \ acc r" + apply (erule rtrancl_induct) + apply blast + apply (blast dest: acc_downward) + done + +theorem acc_downwards: "a \ acc r ==> (b, a) \ r^* ==> b \ acc r" + apply (blast dest: acc_downwards_aux) + done + +theorem acc_wfI: "\x. x \ acc r ==> wf r" + apply (rule wfUNIVI) + apply (induct_tac P x rule: acc_induct) + apply blast + apply blast + done + +theorem acc_wfD: "wf r ==> x \ acc r" + apply (erule wf_induct) + apply (rule accI) + apply blast + done + +theorem wf_acc_iff: "wf r = (\x. x \ acc r)" + apply (blast intro: acc_wfI dest: acc_wfD) + done + +end