# HG changeset patch # User haftmann # Date 1253035007 -7200 # Node ID 986cba8c5053007ed20285af5c929ac21eff2f51 # Parent bf6c78d9f94c2e09203532090ba7d7b80617626d# Parent a382876d3290519aac943bb3482ed43d8421f55c merged diff -r bf6c78d9f94c -r 986cba8c5053 src/HOL/Imperative_HOL/Array.thy --- a/src/HOL/Imperative_HOL/Array.thy Tue Sep 15 15:44:57 2009 +0200 +++ b/src/HOL/Imperative_HOL/Array.thy Tue Sep 15 19:16:47 2009 +0200 @@ -176,12 +176,11 @@ code_type array (OCaml "_/ array") code_const Array (OCaml "failwith/ \"bare Array\"") -code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ _/ _)") +code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ (Big'_int.int'_of'_big'_int/ _)/ _)") code_const Array.of_list (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)") -code_const Array.make' (OCaml "(fun/ ()/ ->/ Array.init/ _/ _)") -code_const Array.length' (OCaml "(fun/ ()/ ->/ Array.length/ _)") -code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ _)") -code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ _/ _)") +code_const Array.length' (OCaml "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))") +code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ (Big'_int.int'_of'_big'_int/ _))") +code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ (Big'_int.int'_of'_big'_int/ _)/ _)") code_reserved OCaml Array diff -r bf6c78d9f94c -r 986cba8c5053 src/HOL/Imperative_HOL/ex/Sublist.thy --- a/src/HOL/Imperative_HOL/ex/Sublist.thy Tue Sep 15 15:44:57 2009 +0200 +++ b/src/HOL/Imperative_HOL/ex/Sublist.thy Tue Sep 15 19:16:47 2009 +0200 @@ -1,4 +1,3 @@ -(* $Id$ *) header {* Slices of lists *} @@ -6,7 +5,6 @@ imports Multiset begin - lemma sublist_split: "i \ j \ j \ k \ sublist xs {i.. ('a \ 'b pred) \ 'b pred" (infixl "\=" 70) where "P \= f = Pred (\x. (\y. eval P y \ eval (f y) x))" -instantiation pred :: (type) complete_lattice +instantiation pred :: (type) "{complete_lattice, boolean_algebra}" begin definition @@ -393,10 +393,16 @@ definition [code del]: "\A = Pred (SUPR A eval)" -instance by default - (auto simp add: less_eq_pred_def less_pred_def +definition + "- P = Pred (- eval P)" + +definition + "P - Q = Pred (eval P - eval Q)" + +instance proof +qed (auto simp add: less_eq_pred_def less_pred_def inf_pred_def sup_pred_def bot_pred_def top_pred_def - Inf_pred_def Sup_pred_def, + Inf_pred_def Sup_pred_def uminus_pred_def minus_pred_def fun_Compl_def bool_Compl_def, auto simp add: le_fun_def less_fun_def le_bool_def less_bool_def eval_inject mem_def) @@ -464,6 +470,127 @@ lemma supE: "eval (A \ B) x \ (eval A x \ P) \ (eval B x \ P) \ P" unfolding sup_pred_def by auto +lemma single_not_bot [simp]: + "single x \ \" + by (auto simp add: single_def bot_pred_def expand_fun_eq) + +lemma not_bot: + assumes "A \ \" + obtains x where "eval A x" +using assms by (cases A) + (auto simp add: bot_pred_def, auto simp add: mem_def) + + +subsubsection {* Emptiness check and definite choice *} + +definition is_empty :: "'a pred \ bool" where + "is_empty A \ A = \" + +lemma is_empty_bot: + "is_empty \" + by (simp add: is_empty_def) + +lemma not_is_empty_single: + "\ is_empty (single x)" + by (auto simp add: is_empty_def single_def bot_pred_def expand_fun_eq) + +lemma is_empty_sup: + "is_empty (A \ B) \ is_empty A \ is_empty B" + by (auto simp add: is_empty_def intro: sup_eq_bot_eq1 sup_eq_bot_eq2) + +definition singleton :: "'a pred \ 'a" where + "singleton A = (if \!x. eval A x then THE x. eval A x else undefined)" + +lemma singleton_eqI: + "\!x. eval A x \ eval A x \ singleton A = x" + by (auto simp add: singleton_def) + +lemma eval_singletonI: + "\!x. eval A x \ eval A (singleton A)" +proof - + assume assm: "\!x. eval A x" + then obtain x where "eval A x" .. + moreover with assm have "singleton A = x" by (rule singleton_eqI) + ultimately show ?thesis by simp +qed + +lemma single_singleton: + "\!x. eval A x \ single (singleton A) = A" +proof - + assume assm: "\!x. eval A x" + then have "eval A (singleton A)" + by (rule eval_singletonI) + moreover from assm have "\x. eval A x \ singleton A = x" + by (rule singleton_eqI) + ultimately have "eval (single (singleton A)) = eval A" + by (simp (no_asm_use) add: single_def expand_fun_eq) blast + then show ?thesis by (simp add: eval_inject) +qed + +lemma singleton_undefinedI: + "\ (\!x. eval A x) \ singleton A = undefined" + by (simp add: singleton_def) + +lemma singleton_bot: + "singleton \ = undefined" + by (auto simp add: bot_pred_def intro: singleton_undefinedI) + +lemma singleton_single: + "singleton (single x) = x" + by (auto simp add: intro: singleton_eqI singleI elim: singleE) + +lemma singleton_sup_single_single: + "singleton (single x \ single y) = (if x = y then x else undefined)" +proof (cases "x = y") + case True then show ?thesis by (simp add: singleton_single) +next + case False + have "eval (single x \ single y) x" + and "eval (single x \ single y) y" + by (auto intro: supI1 supI2 singleI) + with False have "\ (\!z. eval (single x \ single y) z)" + by blast + then have "singleton (single x \ single y) = undefined" + by (rule singleton_undefinedI) + with False show ?thesis by simp +qed + +lemma singleton_sup_aux: + "singleton (A \ B) = (if A = \ then singleton B + else if B = \ then singleton A + else singleton + (single (singleton A) \ single (singleton B)))" +proof (cases "(\!x. eval A x) \ (\!y. eval B y)") + case True then show ?thesis by (simp add: single_singleton) +next + case False + from False have A_or_B: + "singleton A = undefined \ singleton B = undefined" + by (auto intro!: singleton_undefinedI) + then have rhs: "singleton + (single (singleton A) \ single (singleton B)) = undefined" + by (auto simp add: singleton_sup_single_single singleton_single) + from False have not_unique: + "\ (\!x. eval A x) \ \ (\!y. eval B y)" by simp + show ?thesis proof (cases "A \ \ \ B \ \") + case True + then obtain a b where a: "eval A a" and b: "eval B b" + by (blast elim: not_bot) + with True not_unique have "\ (\!x. eval (A \ B) x)" + by (auto simp add: sup_pred_def bot_pred_def) + then have "singleton (A \ B) = undefined" by (rule singleton_undefinedI) + with True rhs show ?thesis by simp + next + case False then show ?thesis by auto + qed +qed + +lemma singleton_sup: + "singleton (A \ B) = (if A = \ then singleton B + else if B = \ then singleton A + else if singleton A = singleton B then singleton A else undefined)" +using singleton_sup_aux [of A B] by (simp only: singleton_sup_single_single) + subsubsection {* Derived operations *} @@ -630,6 +757,46 @@ definition map :: "('a \ 'b) \ 'a pred \ 'b pred" where "map f P = P \= (single o f)" +primrec null :: "'a seq \ bool" where + "null Empty \ True" + | "null (Insert x P) \ False" + | "null (Join P xq) \ is_empty P \ null xq" + +lemma null_is_empty: + "null xq \ is_empty (pred_of_seq xq)" + by (induct xq) (simp_all add: is_empty_bot not_is_empty_single is_empty_sup) + +lemma is_empty_code [code]: + "is_empty (Seq f) \ null (f ())" + by (simp add: null_is_empty Seq_def) + +primrec the_only :: "'a seq \ 'a" where + [code del]: "the_only Empty = undefined" + | "the_only (Insert x P) = (if is_empty P then x else let y = singleton P in if x = y then x else undefined)" + | "the_only (Join P xq) = (if is_empty P then the_only xq else if null xq then singleton P + else let x = singleton P; y = the_only xq in + if x = y then x else undefined)" + +lemma the_only_singleton: + "the_only xq = singleton (pred_of_seq xq)" + by (induct xq) + (auto simp add: singleton_bot singleton_single is_empty_def + null_is_empty Let_def singleton_sup) + +lemma singleton_code [code]: + "singleton (Seq f) = (case f () + of Empty \ undefined + | Insert x P \ if is_empty P then x + else let y = singleton P in + if x = y then x else undefined + | Join P xq \ if is_empty P then the_only xq + else if null xq then singleton P + else let x = singleton P; y = the_only xq in + if x = y then x else undefined)" + by (cases "f ()") + (auto simp add: Seq_def the_only_singleton is_empty_def + null_is_empty singleton_bot singleton_single singleton_sup Let_def) + ML {* signature PREDICATE = sig @@ -707,7 +874,7 @@ bind (infixl "\=" 70) hide (open) type pred seq -hide (open) const Pred eval single bind if_pred not_pred - Empty Insert Join Seq member pred_of_seq "apply" adjunct eq map +hide (open) const Pred eval single bind is_empty singleton if_pred not_pred + Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map end diff -r bf6c78d9f94c -r 986cba8c5053 src/HOL/ex/Predicate_Compile_ex.thy --- a/src/HOL/ex/Predicate_Compile_ex.thy Tue Sep 15 15:44:57 2009 +0200 +++ b/src/HOL/ex/Predicate_Compile_ex.thy Tue Sep 15 19:16:47 2009 +0200 @@ -157,4 +157,13 @@ values 3 "{(a,q). step (par nil nil) a q}" *) + +inductive divmod_rel :: "nat \ nat \ nat \ nat \ bool" where + "k < l \ divmod_rel k l 0 k" + | "k \ l \ divmod_rel (k - l) l q r \ divmod_rel k l (Suc q) r" + +code_pred divmod_rel .. + +value [code] "Predicate.singleton (divmod_rel_1_2 1705 42)" + end \ No newline at end of file