# HG changeset patch # User paulson # Date 972377331 -7200 # Node ID ec30a7d15f76c2d8e611ef1c98d914f6f1357196 # Parent 5b36035e4dff598b2ff51e2d13bea9afeaa55b7c Acc example diff -r 5b36035e4dff -r ec30a7d15f76 doc-src/TutorialI/Inductive/Acc.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/TutorialI/Inductive/Acc.thy Tue Oct 24 10:48:51 2000 +0200 @@ -0,0 +1,78 @@ +(* Title: HOL/ex/Acc.thy + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 University of Cambridge + +Inductive definition of acc(r) + +See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. +Research Report 92-49, LIP, ENS Lyon. Dec 1992. +*) + +header {* The accessible part of a relation *} + +theory Acc = Main: + +consts + acc :: "('a \ 'a) set => 'a set" -- {* accessible part *} + +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)" + + +theorem acc_induct: + "[| 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