Up to index of Isabelle/HOL
theory Accessible_Part(* Title: HOL/Accessible_Part.thy ID: $Id: Accessible_Part.thy,v 1.6 2007/07/16 19:17:12 krauss Exp $ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) header {* The accessible part of a relation *} theory Accessible_Part imports Wellfounded_Recursion begin subsection {* Inductive definition *} text {* Inductive definition of the accessible part @{term "acc r"} of a relation; see also \cite{paulin-tlca}. *} inductive_set acc :: "('a * 'a) set => 'a set" for r :: "('a * 'a) set" where accI: "(!!y. (y, x) : r ==> y : acc r) ==> x : acc r" abbreviation termip :: "('a => 'a => bool) => 'a => bool" where "termip r == accp (r¯¯)" abbreviation termi :: "('a * 'a) set => 'a set" where "termi r == acc (r¯)" lemmas accpI = accp.accI subsection {* Induction rules *} theorem accp_induct: assumes major: "accp r a" assumes hyp: "!!x. accp r x ==> ∀y. r y x --> P y ==> P x" shows "P a" apply (rule major [THEN accp.induct]) apply (rule hyp) apply (rule accp.accI) apply fast apply fast done theorems accp_induct_rule = accp_induct [rule_format, induct set: accp] theorem accp_downward: "accp r b ==> r a b ==> accp r a" apply (erule accp.cases) apply fast done lemma not_accp_down: assumes na: "¬ accp R x" obtains z where "R z x" and "¬ accp R z" proof - assume a: "!!z. [|R z x; ¬ accp R z|] ==> thesis" show thesis proof (cases "∀z. R z x --> accp R z") case True hence "!!z. R z x ==> accp R z" by auto hence "accp R x" by (rule accp.accI) with na show thesis .. next case False then obtain z where "R z x" and "¬ accp R z" by auto with a show thesis . qed qed lemma accp_downwards_aux: "r** b a ==> accp r a --> accp r b" apply (erule rtranclp_induct) apply blast apply (blast dest: accp_downward) done theorem accp_downwards: "accp r a ==> r** b a ==> accp r b" apply (blast dest: accp_downwards_aux) done theorem accp_wfPI: "∀x. accp r x ==> wfP r" apply (rule wfPUNIVI) apply (induct_tac P x rule: accp_induct) apply blast apply blast done theorem accp_wfPD: "wfP r ==> accp r x" apply (erule wfP_induct_rule) apply (rule accp.accI) apply blast done theorem wfP_accp_iff: "wfP r = (∀x. accp r x)" apply (blast intro: accp_wfPI dest: accp_wfPD) done text {* Smaller relations have bigger accessible parts: *} lemma accp_subset: assumes sub: "R1 ≤ R2" shows "accp R2 ≤ accp R1" proof fix x assume "accp R2 x" then show "accp R1 x" proof (induct x) fix x assume ih: "!!y. R2 y x ==> accp R1 y" with sub show "accp R1 x" by (blast intro: accp.accI) qed qed text {* This is a generalized induction theorem that works on subsets of the accessible part. *} lemma accp_subset_induct: assumes subset: "D ≤ accp R" and dcl: "!!x z. [|D x; R z x|] ==> D z" and "D x" and istep: "!!x. [|D x; (!!z. R z x ==> P z)|] ==> P x" shows "P x" proof - from subset and `D x` have "accp R x" .. then show "P x" using `D x` proof (induct x) fix x assume "D x" and "!!y. R y x ==> D y ==> P y" with dcl and istep show "P x" by blast qed qed text {* Set versions of the above theorems *} lemmas acc_induct = accp_induct [to_set] lemmas acc_induct_rule = acc_induct [rule_format, induct set: acc] lemmas acc_downward = accp_downward [to_set] lemmas not_acc_down = not_accp_down [to_set] lemmas acc_downwards_aux = accp_downwards_aux [to_set] lemmas acc_downwards = accp_downwards [to_set] lemmas acc_wfI = accp_wfPI [to_set] lemmas acc_wfD = accp_wfPD [to_set] lemmas wf_acc_iff = wfP_accp_iff [to_set] lemmas acc_subset = accp_subset [to_set] lemmas acc_subset_induct = accp_subset_induct [to_set] end
lemma accpI:
(!!y. r y x ==> accp r y) ==> accp r x
theorem accp_induct:
[| accp r a; !!x. [| accp r x; ∀y. r y x --> P y |] ==> P x |] ==> P a
theorem accp_induct_rule:
[| accp r a; !!x. [| accp r x; !!y. r y x ==> P y |] ==> P x |] ==> P a
theorem accp_downward:
[| accp r b; r a b |] ==> accp r a
lemma not_accp_down:
[| ¬ accp R x; !!z. [| R z x; ¬ accp R z |] ==> thesis |] ==> thesis
lemma accp_downwards_aux:
r** b a ==> accp r a --> accp r b
theorem accp_downwards:
[| accp r a; r** b a |] ==> accp r b
theorem accp_wfPI:
∀x. accp r x ==> wfP r
theorem accp_wfPD:
wfP r ==> accp r x
theorem wfP_accp_iff:
wfP r = (∀x. accp r x)
lemma accp_subset:
R1.0 ≤ R2.0 ==> accp R2.0 ≤ accp R1.0
lemma accp_subset_induct:
[| D ≤ accp R; !!x z. [| D x; R z x |] ==> D z; D x;
!!x. [| D x; !!z. R z x ==> P z |] ==> P x |]
==> P x
lemma acc_induct:
[| a ∈ acc r; !!x. [| x ∈ acc r; ∀y. (y, x) ∈ r --> P y |] ==> P x |] ==> P a
lemma acc_induct_rule:
[| a ∈ acc r; !!x. [| x ∈ acc r; !!y. (y, x) ∈ r ==> P y |] ==> P x |] ==> P a
lemma acc_downward:
[| b ∈ acc r; (a, b) ∈ r |] ==> a ∈ acc r
lemma not_acc_down:
[| x ∉ acc R; !!z. [| (z, x) ∈ R; z ∉ acc R |] ==> thesis |] ==> thesis
lemma acc_downwards_aux:
(b, a) ∈ r* ==> a ∈ acc r --> b ∈ acc r
lemma acc_downwards:
[| a ∈ acc r; (b, a) ∈ r* |] ==> b ∈ acc r
lemma acc_wfI:
∀x. x ∈ acc r ==> wf r
lemma acc_wfD:
wf r ==> x ∈ acc r
lemma wf_acc_iff:
wf r = (∀x. x ∈ acc r)
lemma acc_subset:
R1.0 ⊆ R2.0 ==> acc R2.0 ⊆ acc R1.0
lemma acc_subset_induct:
[| D ⊆ acc R; !!x z. [| x ∈ D; (z, x) ∈ R |] ==> z ∈ D; x ∈ D;
!!x. [| x ∈ D; !!z. (z, x) ∈ R ==> P z |] ==> P x |]
==> P x