# HG changeset patch # User nipkow # Date 1371569871 -7200 # Node ID 3971dd9ca831ed1223e44b8e161aedf1263eee3f # Parent b5b510c686cb32c46433fa2446539ef2b47e8650 Added continuity and determinism proof diff -r b5b510c686cb -r 3971dd9ca831 src/HOL/IMP/Denotation.thy --- a/src/HOL/IMP/Denotation.thy Tue Jun 18 10:04:06 2013 +0200 +++ b/src/HOL/IMP/Denotation.thy Tue Jun 18 17:37:51 2013 +0200 @@ -9,51 +9,153 @@ definition W :: "bexp \ com_den \ (com_den \ com_den)" where "W b rc = (\rw. {(s,t). if bval b s then (s,t) \ rc O rw else s=t})" -fun C :: "com \ com_den" where -"C SKIP = {(s,t). s=t}" | -"C (x ::= a) = {(s,t). t = s(x := aval a s)}" | -"C (c0;;c1) = C(c0) O C(c1)" | -"C (IF b THEN c1 ELSE c2) - = {(s,t). if bval b s then (s,t) \ C c1 else (s,t) \ C c2}" | -"C(WHILE b DO c) = lfp (W b (C c))" +fun D :: "com \ com_den" where +"D SKIP = Id" | +"D (x ::= a) = {(s,t). t = s(x := aval a s)}" | +"D (c0;;c1) = D(c0) O D(c1)" | +"D (IF b THEN c1 ELSE c2) + = {(s,t). if bval b s then (s,t) \ D c1 else (s,t) \ D c2}" | +"D (WHILE b DO c) = lfp (W b (D c))" lemma W_mono: "mono (W b r)" by (unfold W_def mono_def) auto -lemma C_While_If: - "C(WHILE b DO c) = C(IF b THEN c;;WHILE b DO c ELSE SKIP)" -apply simp -apply (subst lfp_unfold [OF W_mono]) --{*lhs only*} -apply (simp add: W_def) -done +lemma D_While_If: + "D(WHILE b DO c) = D(IF b THEN c;;WHILE b DO c ELSE SKIP)" +proof- + let ?w = "WHILE b DO c" + have "D ?w = lfp (W b (D c))" by simp + also have "\ = W b (D c) (lfp (W b (D c)))" by(rule lfp_unfold [OF W_mono]) + also have "\ = D(IF b THEN c;;?w ELSE SKIP)" by (simp add: W_def) + finally show ?thesis . +qed text{* Equivalence of denotational and big-step semantics: *} -lemma C_if_big_step: "(c,s) \ t \ (s,t) \ C(c)" +lemma D_if_big_step: "(c,s) \ t \ (s,t) \ D(c)" proof (induction rule: big_step_induct) case WhileFalse - with C_While_If show ?case by auto + with D_While_If show ?case by auto next case WhileTrue - show ?case unfolding C_While_If using WhileTrue by auto + show ?case unfolding D_While_If using WhileTrue by auto qed auto abbreviation Big_step :: "com \ com_den" where "Big_step c \ {(s,t). (c,s) \ t}" -lemma Big_step_if_C: "(s,t) \ C(c) \ (s,t) : Big_step c" +lemma Big_step_if_D: "(s,t) \ D(c) \ (s,t) : Big_step c" proof (induction c arbitrary: s t) case Seq thus ?case by fastforce next case (While b c) let ?B = "Big_step (WHILE b DO c)" - have "W b (C c) ?B <= ?B" using While.IH by (auto simp: W_def) - from lfp_lowerbound[where ?f = "W b (C c)", OF this] While.prems + have "W b (D c) ?B <= ?B" using While.IH by (auto simp: W_def) + from lfp_lowerbound[where ?f = "W b (D c)", OF this] While.prems show ?case by auto qed (auto split: if_splits) theorem denotational_is_big_step: - "(s,t) \ C(c) = ((c,s) \ t)" -by (metis C_if_big_step Big_step_if_C[simplified]) + "(s,t) \ D(c) = ((c,s) \ t)" +by (metis D_if_big_step Big_step_if_D[simplified]) + + +subsection "Continuity" + +definition chain :: "(nat \ 'a set) \ bool" where +"chain S = (\i. S i \ S(Suc i))" + +lemma chain_total: "chain S \ S i \ S j \ S j \ S i" +by (metis chain_def le_cases lift_Suc_mono_le) + +definition cont :: "('a set \ 'a set) \ bool" where +"cont f = (\S. chain S \ f(UN n. S n) = (UN n. f(S n)))" + +lemma mono_if_cont: fixes f :: "'a set \ 'a set" + assumes "cont f" shows "mono f" +proof + fix a b :: "'a set" assume "a \ b" + let ?S = "\n::nat. if n=0 then a else b" + have "chain ?S" using `a \ b` by(auto simp: chain_def) + hence "f(UN n. ?S n) = (UN n. f(?S n))" using assms by(simp add: cont_def) + moreover have "(UN n. ?S n) = b" using `a \ b` by (auto split: if_splits) + moreover have "(UN n. f(?S n)) = f a \ f b" by (auto split: if_splits) + ultimately show "f a \ f b" by (metis Un_upper1) +qed + +lemma chain_iterates: fixes f :: "'a set \ 'a set" + assumes "mono f" shows "chain(\n. (f^^n) {})" +proof- + { fix n have "(f ^^ n) {} \ (f ^^ Suc n) {}" using assms + by(induction n) (auto simp: mono_def) } + thus ?thesis by(auto simp: chain_def) +qed + +theorem lfp_if_cont: + assumes "cont f" shows "lfp f = (UN n. (f^^n) {})" (is "_ = ?U") +proof + show "lfp f \ ?U" + proof (rule lfp_lowerbound) + have "f ?U = (UN n. (f^^Suc n){})" + using chain_iterates[OF mono_if_cont[OF assms]] assms + by(simp add: cont_def) + also have "\ = (f^^0){} \ \" by simp + also have "\ = ?U" + by(auto simp del: funpow.simps) (metis not0_implies_Suc) + finally show "f ?U \ ?U" by simp + qed +next + { fix n p assume "f p \ p" + have "(f^^n){} \ p" + proof(induction n) + case 0 show ?case by simp + next + case Suc + from monoD[OF mono_if_cont[OF assms] Suc] `f p \ p` + show ?case by simp + qed + } + thus "?U \ lfp f" by(auto simp: lfp_def) +qed + +lemma cont_W: "cont(W b r)" +by(auto simp: cont_def W_def) + +subsection{*The denotational semantics is deterministic*} + +(* FIXME mv *) +lemma simgle_valued_empty[simp]: "single_valued {}" +by(simp add: single_valued_def) + +lemma single_valued_UN_chain: + assumes "chain S" "(\n. single_valued (S n))" + shows "single_valued(UN n. S n)" +proof(auto simp: single_valued_def) + fix m n x y z assume "(x, y) \ S m" "(x, z) \ S n" + with chain_total[OF assms(1), of m n] assms(2) + show "y = z" by (auto simp: single_valued_def) +qed + +lemma single_valued_lfp: fixes f :: "com_den \ com_den" +assumes "cont f" "\r. single_valued r \ single_valued (f r)" +shows "single_valued(lfp f)" +unfolding lfp_if_cont[OF assms(1)] +proof(rule single_valued_UN_chain[OF chain_iterates[OF mono_if_cont[OF assms(1)]]]) + fix n show "single_valued ((f ^^ n) {})" + by(induction n)(auto simp: assms(2)) +qed + +lemma single_valued_D: "single_valued (D c)" +proof(induction c) + case Seq thus ?case by(simp add: single_valued_relcomp) +next + case (While b c) + have "single_valued (lfp (W b (D c)))" + proof(rule single_valued_lfp[OF cont_W]) + show "!!r. single_valued r \ single_valued (W b (D c) r)" + using While.IH by(force simp: single_valued_def W_def) + qed + thus ?case by simp +qed (auto simp add: single_valued_def) end