oheimb@11351: (* Title: HOL/Library/Continuity.thy wenzelm@11355: ID: $Id$ wenzelm@11355: Author: David von Oheimb, TU Muenchen oheimb@11351: License: GPL (GNU GENERAL PUBLIC LICENSE) oheimb@11351: *) oheimb@11351: oheimb@11351: header {* wenzelm@11355: \title{Continuity and iterations (of set transformers)} oheimb@11351: \author{David von Oheimb} oheimb@11351: *} oheimb@11351: wenzelm@11355: theory Continuity = Main: oheimb@11351: oheimb@11351: subsection "Chains" oheimb@11351: oheimb@11351: constdefs wenzelm@11355: up_chain :: "(nat => 'a set) => bool" wenzelm@11355: "up_chain F == \i. F i \ F (Suc i)" oheimb@11351: wenzelm@11355: lemma up_chainI: "(!!i. F i \ F (Suc i)) ==> up_chain F" wenzelm@11355: by (simp add: up_chain_def) oheimb@11351: wenzelm@11355: lemma up_chainD: "up_chain F ==> F i \ F (Suc i)" wenzelm@11355: by (simp add: up_chain_def) oheimb@11351: wenzelm@11355: lemma up_chain_less_mono [rule_format]: wenzelm@11355: "up_chain F ==> x < y --> F x \ F y" wenzelm@11355: apply (induct_tac y) wenzelm@11355: apply (blast dest: up_chainD elim: less_SucE)+ wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma up_chain_mono: "up_chain F ==> x \ y ==> F x \ F y" wenzelm@11355: apply (drule le_imp_less_or_eq) wenzelm@11355: apply (blast dest: up_chain_less_mono) wenzelm@11355: done oheimb@11351: oheimb@11351: oheimb@11351: constdefs wenzelm@11355: down_chain :: "(nat => 'a set) => bool" wenzelm@11355: "down_chain F == \i. F (Suc i) \ F i" oheimb@11351: wenzelm@11355: lemma down_chainI: "(!!i. F (Suc i) \ F i) ==> down_chain F" wenzelm@11355: by (simp add: down_chain_def) oheimb@11351: wenzelm@11355: lemma down_chainD: "down_chain F ==> F (Suc i) \ F i" wenzelm@11355: by (simp add: down_chain_def) oheimb@11351: wenzelm@11355: lemma down_chain_less_mono [rule_format]: wenzelm@11355: "down_chain F ==> x < y --> F y \ F x" wenzelm@11355: apply (induct_tac y) wenzelm@11355: apply (blast dest: down_chainD elim: less_SucE)+ wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma down_chain_mono: "down_chain F ==> x \ y ==> F y \ F x" wenzelm@11355: apply (drule le_imp_less_or_eq) wenzelm@11355: apply (blast dest: down_chain_less_mono) wenzelm@11355: done oheimb@11351: oheimb@11351: oheimb@11351: subsection "Continuity" oheimb@11351: oheimb@11351: constdefs oheimb@11351: up_cont :: "('a set => 'a set) => bool" wenzelm@11355: "up_cont f == \F. up_chain F --> f (\(range F)) = \(f ` range F)" oheimb@11351: wenzelm@11355: lemma up_contI: wenzelm@11355: "(!!F. up_chain F ==> f (\(range F)) = \(f ` range F)) ==> up_cont f" wenzelm@11355: apply (unfold up_cont_def) wenzelm@11355: apply blast wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma up_contD: wenzelm@11355: "up_cont f ==> up_chain F ==> f (\(range F)) = \(f ` range F)" wenzelm@11355: apply (unfold up_cont_def) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: oheimb@11351: oheimb@11351: lemma up_cont_mono: "up_cont f ==> mono f" wenzelm@11355: apply (rule monoI) wenzelm@11355: apply (drule_tac F = "\i. if i = 0 then A else B" in up_contD) wenzelm@11355: apply (rule up_chainI) wenzelm@11355: apply simp+ wenzelm@11355: apply (drule Un_absorb1) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: oheimb@11351: oheimb@11351: constdefs oheimb@11351: down_cont :: "('a set => 'a set) => bool" wenzelm@11355: "down_cont f == wenzelm@11355: \F. down_chain F --> f (Inter (range F)) = Inter (f ` range F)" oheimb@11351: wenzelm@11355: lemma down_contI: wenzelm@11355: "(!!F. down_chain F ==> f (Inter (range F)) = Inter (f ` range F)) ==> wenzelm@11355: down_cont f" wenzelm@11355: apply (unfold down_cont_def) wenzelm@11355: apply blast wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma down_contD: "down_cont f ==> down_chain F ==> wenzelm@11355: f (Inter (range F)) = Inter (f ` range F)" wenzelm@11355: apply (unfold down_cont_def) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: oheimb@11351: lemma down_cont_mono: "down_cont f ==> mono f" wenzelm@11355: apply (rule monoI) wenzelm@11355: apply (drule_tac F = "\i. if i = 0 then B else A" in down_contD) wenzelm@11355: apply (rule down_chainI) wenzelm@11355: apply simp+ wenzelm@11355: apply (drule Int_absorb1) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: oheimb@11351: oheimb@11351: subsection "Iteration" oheimb@11351: oheimb@11351: constdefs oheimb@11351: up_iterate :: "('a set => 'a set) => nat => 'a set" wenzelm@11355: "up_iterate f n == (f^n) {}" oheimb@11351: oheimb@11351: lemma up_iterate_0 [simp]: "up_iterate f 0 = {}" wenzelm@11355: by (simp add: up_iterate_def) oheimb@11351: wenzelm@11355: lemma up_iterate_Suc [simp]: "up_iterate f (Suc i) = f (up_iterate f i)" wenzelm@11355: by (simp add: up_iterate_def) oheimb@11351: oheimb@11351: lemma up_iterate_chain: "mono F ==> up_chain (up_iterate F)" wenzelm@11355: apply (rule up_chainI) wenzelm@11355: apply (induct_tac i) wenzelm@11355: apply simp+ wenzelm@11355: apply (erule (1) monoD) wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma UNION_up_iterate_is_fp: wenzelm@11355: "up_cont F ==> wenzelm@11355: F (UNION UNIV (up_iterate F)) = UNION UNIV (up_iterate F)" wenzelm@11355: apply (frule up_cont_mono [THEN up_iterate_chain]) wenzelm@11355: apply (drule (1) up_contD) wenzelm@11355: apply simp wenzelm@11355: apply (auto simp del: up_iterate_Suc simp add: up_iterate_Suc [symmetric]) wenzelm@11355: apply (case_tac xa) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma UNION_up_iterate_lowerbound: wenzelm@11355: "mono F ==> F P = P ==> UNION UNIV (up_iterate F) \ P" wenzelm@11355: apply (subgoal_tac "(!!i. up_iterate F i \ P)") wenzelm@11355: apply fast wenzelm@11355: apply (induct_tac i) wenzelm@11355: prefer 2 apply (drule (1) monoD) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma UNION_up_iterate_is_lfp: wenzelm@11355: "up_cont F ==> lfp F = UNION UNIV (up_iterate F)" wenzelm@11355: apply (rule set_eq_subset [THEN iffD2]) wenzelm@11355: apply (rule conjI) wenzelm@11355: prefer 2 wenzelm@11355: apply (drule up_cont_mono) wenzelm@11355: apply (rule UNION_up_iterate_lowerbound) wenzelm@11355: apply assumption wenzelm@11355: apply (erule lfp_unfold [symmetric]) wenzelm@11355: apply (rule lfp_lowerbound) wenzelm@11355: apply (rule set_eq_subset [THEN iffD1, THEN conjunct2]) wenzelm@11355: apply (erule UNION_up_iterate_is_fp [symmetric]) wenzelm@11355: done oheimb@11351: oheimb@11351: oheimb@11351: constdefs oheimb@11351: down_iterate :: "('a set => 'a set) => nat => 'a set" wenzelm@11355: "down_iterate f n == (f^n) UNIV" oheimb@11351: oheimb@11351: lemma down_iterate_0 [simp]: "down_iterate f 0 = UNIV" wenzelm@11355: by (simp add: down_iterate_def) oheimb@11351: wenzelm@11355: lemma down_iterate_Suc [simp]: wenzelm@11355: "down_iterate f (Suc i) = f (down_iterate f i)" wenzelm@11355: by (simp add: down_iterate_def) oheimb@11351: oheimb@11351: lemma down_iterate_chain: "mono F ==> down_chain (down_iterate F)" wenzelm@11355: apply (rule down_chainI) wenzelm@11355: apply (induct_tac i) wenzelm@11355: apply simp+ wenzelm@11355: apply (erule (1) monoD) wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma INTER_down_iterate_is_fp: wenzelm@11355: "down_cont F ==> wenzelm@11355: F (INTER UNIV (down_iterate F)) = INTER UNIV (down_iterate F)" wenzelm@11355: apply (frule down_cont_mono [THEN down_iterate_chain]) wenzelm@11355: apply (drule (1) down_contD) wenzelm@11355: apply simp wenzelm@11355: apply (auto simp del: down_iterate_Suc simp add: down_iterate_Suc [symmetric]) wenzelm@11355: apply (case_tac xa) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma INTER_down_iterate_upperbound: wenzelm@11355: "mono F ==> F P = P ==> P \ INTER UNIV (down_iterate F)" wenzelm@11355: apply (subgoal_tac "(!!i. P \ down_iterate F i)") wenzelm@11355: apply fast wenzelm@11355: apply (induct_tac i) wenzelm@11355: prefer 2 apply (drule (1) monoD) wenzelm@11355: apply auto wenzelm@11355: done oheimb@11351: wenzelm@11355: lemma INTER_down_iterate_is_gfp: wenzelm@11355: "down_cont F ==> gfp F = INTER UNIV (down_iterate F)" wenzelm@11355: apply (rule set_eq_subset [THEN iffD2]) wenzelm@11355: apply (rule conjI) wenzelm@11355: apply (drule down_cont_mono) wenzelm@11355: apply (rule INTER_down_iterate_upperbound) wenzelm@11355: apply assumption wenzelm@11355: apply (erule gfp_unfold [symmetric]) wenzelm@11355: apply (rule gfp_upperbound) wenzelm@11355: apply (rule set_eq_subset [THEN iffD1, THEN conjunct2]) wenzelm@11355: apply (erule INTER_down_iterate_is_fp) wenzelm@11355: done oheimb@11351: oheimb@11351: end