# HG changeset patch # User Christian Urban # Date 1401717522 -3600 # Node ID 7ea6b019ce24fe84c3b54ee26163aed8774c9936 # Parent c89013dca1aa71df7907631f3c8812de061dea14 updated diff -r c89013dca1aa -r 7ea6b019ce24 PrioG.thy --- a/PrioG.thy Fri May 30 07:56:39 2014 +0100 +++ b/PrioG.thy Mon Jun 02 14:58:42 2014 +0100 @@ -804,19 +804,6 @@ apply(auto) done -(* FIXME: Unused -lemma simple_A: - fixes A - assumes h: "\ x y. \x \ A; y \ A\ \ x = y" - shows "A = {} \ (\ a. A = {a})" -proof(cases "A = {}") - case True thus ?thesis by simp -next - case False then obtain a where "a \ A" by auto - with h have "A = {a}" by auto - thus ?thesis by simp -qed -*) lemma RAG_target_th: "(Th th, x) \ RAG (s::state) \ \ cs. x = Cs cs" by (unfold s_RAG_def, auto) @@ -2092,7 +2079,7 @@ apply(simp add: Let_def) done - +(* FIXME: NOT NEEDED *) lemma runing_unique: fixes th1 th2 s assumes vt: "vt s" @@ -2281,6 +2268,7 @@ qed qed + lemma create_pre: assumes stp: "step s e" and not_in: "th \ threads s" diff -r c89013dca1aa -r 7ea6b019ce24 ROOT --- a/ROOT Fri May 30 07:56:39 2014 +0100 +++ b/ROOT Mon Jun 02 14:58:42 2014 +0100 @@ -1,7 +1,8 @@ session "PIP" = HOL + theories [document = false] "CpsG" - "ExtGG" + "ExtGG" + "Test" session "Slides2" in "Slides" = PIP + options [document_variants="slides2"] diff -r c89013dca1aa -r 7ea6b019ce24 Test.thy --- a/Test.thy Fri May 30 07:56:39 2014 +0100 +++ b/Test.thy Mon Jun 02 14:58:42 2014 +0100 @@ -350,7 +350,70 @@ qed (simp_all) +lemma finite_RAG: + assumes "vt s" + shows "finite (RAG2 s)" +using assms +apply(induct) +apply(simp add: RAG2_def RAG_def waits_def holds_def) +apply(erule step.cases) +apply(auto) +apply(case_tac "wq sa cs = []") +apply(rule finite_subset) +apply(rule RAG_p1) +apply(simp) +apply(simp) +apply(rule finite_subset) +apply(rule RAG_p2) +apply(simp) +apply(simp) +apply(simp) +apply(subgoal_tac "\wts. wq sa cs = th # wts") +apply(erule exE) +apply(case_tac "wts = []") +apply(rule finite_subset) +apply(rule RAG_v1) +apply(simp) +apply(simp) +apply(rule finite_subset) +apply(rule RAG_v2) +apply(simp) +apply(simp) +apply(simp) +apply(subgoal_tac "th \ set (wq sa cs) \ th = hd (wq sa cs)") +apply(case_tac "wq sa cs") +apply(auto)[2] +apply(auto simp add: holds2_def holds_def wq_def) +done + +lemma wq_threads: + assumes vt: "vt s" + and h: "th \ set (wq s cs)" + shows "th \ threads s" +using assms +apply(induct) +apply(simp add: wq_def) +apply(erule step.cases) +apply(auto simp add: wq_def Let_def holding_def holds2_def holds_def waits2_def runing_def readys_def) +apply(simp add: waits_def) +apply(auto simp add: waits_def split: if_splits)[1] +apply(auto split: if_splits) +apply(simp only: waits_def) +by (metis insert_iff set_simps(2)) + +lemma max_cpreceds_readys_threads: + assumes vt: "vt s" + shows "Max (cpreceds2 s (readys s)) = Max (cpreceds2 s (threads s))" +apply(case_tac "threads s = {}") +apply(simp add: readys_def) +using vt +apply(induct) +apply(simp) +apply(erule step.cases) +apply(auto simp add: readys_def waits2_def waits_def Let_def holding_def runing_def holds2_def) +defer +defer +oops - - +end \ No newline at end of file diff -r c89013dca1aa -r 7ea6b019ce24 journal.pdf Binary file journal.pdf has changed