proper merge of operation changes and generic do-syntax
authorhaftmann
Tue, 13 Jul 2010 12:16:24 +0200
changeset 37798 0b0570445a2a
parent 37797 96551d6b1414
child 37799 b2f84bb86c73
proper merge of operation changes and generic do-syntax
src/HOL/Imperative_HOL/Array.thy
src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy
src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy
src/HOL/Imperative_HOL/ex/SatChecker.thy
--- a/src/HOL/Imperative_HOL/Array.thy	Tue Jul 13 12:05:20 2010 +0200
+++ b/src/HOL/Imperative_HOL/Array.thy	Tue Jul 13 12:16:24 2010 +0200
@@ -35,9 +35,8 @@
   length :: "'a\<Colon>heap array \<Rightarrow> heap \<Rightarrow> nat" where
   "length a h = List.length (get_array a h)"
   
-definition (*FIXME update*)
-  change :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
-  "change a i x h = set_array a ((get_array a h)[i:=x]) h"
+definition update :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
+  "update a i x h = set_array a ((get_array a h)[i:=x]) h"
 
 definition (*FIXME noteq*)
   noteq_arrs :: "'a\<Colon>heap array \<Rightarrow> 'b\<Colon>heap array \<Rightarrow> bool" (infix "=!!=" 70) where
@@ -64,15 +63,15 @@
 
 definition upd :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a\<Colon>heap array Heap" where
   [code del]: "upd i x a = Heap_Monad.guard (\<lambda>h. i < length a h)
-    (\<lambda>h. (a, change a i x h))"
+    (\<lambda>h. (a, update a i x h))"
 
 definition map_entry :: "nat \<Rightarrow> ('a\<Colon>heap \<Rightarrow> 'a) \<Rightarrow> 'a array \<Rightarrow> 'a array Heap" where
   [code del]: "map_entry i f a = Heap_Monad.guard (\<lambda>h. i < length a h)
-    (\<lambda>h. (a, change a i (f (get_array a h ! i)) h))"
+    (\<lambda>h. (a, update a i (f (get_array a h ! i)) h))"
 
 definition swap :: "nat \<Rightarrow> 'a \<Rightarrow> 'a\<Colon>heap array \<Rightarrow> 'a Heap" where
   [code del]: "swap i x a = Heap_Monad.guard (\<lambda>h. i < length a h)
-    (\<lambda>h. (get_array a h ! i, change a i x h))"
+    (\<lambda>h. (get_array a h ! i, update a i x h))"
 
 definition freeze :: "'a\<Colon>heap array \<Rightarrow> 'a list Heap" where
   [code del]: "freeze a = Heap_Monad.tap (\<lambda>h. get_array a h)"
@@ -109,27 +108,27 @@
   "r =!!= r' \<Longrightarrow> set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)"
   by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def)
 
-lemma get_array_change_eq [simp]:
-  "get_array a (change a i v h) = (get_array a h) [i := v]"
-  by (simp add: change_def)
+lemma get_array_update_eq [simp]:
+  "get_array a (update a i v h) = (get_array a h) [i := v]"
+  by (simp add: update_def)
 
-lemma nth_change_array_neq_array [simp]:
-  "a =!!= b \<Longrightarrow> get_array a (change b j v h) ! i = get_array a h ! i"
-  by (simp add: change_def noteq_arrs_def)
+lemma nth_update_array_neq_array [simp]:
+  "a =!!= b \<Longrightarrow> get_array a (update b j v h) ! i = get_array a h ! i"
+  by (simp add: update_def noteq_arrs_def)
 
-lemma get_arry_array_change_elem_neqIndex [simp]:
-  "i \<noteq> j \<Longrightarrow> get_array a (change a j v h) ! i = get_array a h ! i"
+lemma get_arry_array_update_elem_neqIndex [simp]:
+  "i \<noteq> j \<Longrightarrow> get_array a (update a j v h) ! i = get_array a h ! i"
   by simp
 
-lemma length_change [simp]: 
-  "length a (change b i v h) = length a h"
-  by (simp add: change_def length_def set_array_def get_array_def)
+lemma length_update [simp]: 
+  "length a (update b i v h) = length a h"
+  by (simp add: update_def length_def set_array_def get_array_def)
 
-lemma change_swap_neqArray:
+lemma update_swap_neqArray:
   "a =!!= a' \<Longrightarrow> 
-  change a i v (change a' i' v' h) 
-  = change a' i' v' (change a i v h)"
-apply (unfold change_def)
+  update a i v (update a' i' v' h) 
+  = update a' i' v' (update a i v h)"
+apply (unfold update_def)
 apply simp
 apply (subst array_set_set_swap, assumption)
 apply (subst array_get_set_neq)
@@ -137,9 +136,9 @@
 apply (simp)
 done
 
-lemma change_swap_neqIndex:
-  "\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> change a i v (change a i' v' h) = change a i' v' (change a i v h)"
-  by (auto simp add: change_def array_set_set_swap list_update_swap)
+lemma update_swap_neqIndex:
+  "\<lbrakk> i \<noteq> i' \<rbrakk> \<Longrightarrow> update a i v (update a i' v' h) = update a i' v' (update a i v h)"
+  by (auto simp add: update_def array_set_set_swap list_update_swap)
 
 lemma get_array_init_array_list:
   "get_array (fst (array ls h)) (snd (array ls' h)) = ls'"
@@ -151,9 +150,9 @@
        = snd (array new_ls h)"
   by (simp add: Let_def split_def array_def)
 
-lemma array_present_change [simp]: 
-  "array_present a (change b i v h) = array_present a h"
-  by (simp add: change_def array_present_def set_array_def get_array_def)
+lemma array_present_update [simp]: 
+  "array_present a (update b i v h) = array_present a h"
+  by (simp add: update_def array_present_def set_array_def get_array_def)
 
 lemma array_present_array [simp]:
   "array_present (fst (array xs h)) (snd (array xs h))"
@@ -264,7 +263,7 @@
 
 lemma execute_upd [execute_simps]:
   "i < length a h \<Longrightarrow>
-    execute (upd i x a) h = Some (a, change a i x h)"
+    execute (upd i x a) h = Some (a, update a i x h)"
   "i \<ge> length a h \<Longrightarrow> execute (upd i x a) h = None"
   by (simp_all add: upd_def execute_simps)
 
@@ -273,20 +272,20 @@
   by (auto intro: success_intros simp add: upd_def)
 
 lemma crel_updI [crel_intros]:
-  assumes "i < length a h" "h' = change a i v h"
+  assumes "i < length a h" "h' = update a i v h"
   shows "crel (upd i v a) h h' a"
   by (rule crelI) (insert assms, simp add: execute_simps)
 
 lemma crel_updE [crel_elims]:
   assumes "crel (upd i v a) h h' r"
-  obtains "r = a" "h' = change a i v h" "i < length a h"
+  obtains "r = a" "h' = update a i v h" "i < length a h"
   using assms by (rule crelE)
     (erule successE, cases "i < length a h", simp_all add: execute_simps)
 
 lemma execute_map_entry [execute_simps]:
   "i < length a h \<Longrightarrow>
    execute (map_entry i f a) h =
-      Some (a, change a i (f (get_array a h ! i)) h)"
+      Some (a, update a i (f (get_array a h ! i)) h)"
   "i \<ge> length a h \<Longrightarrow> execute (map_entry i f a) h = None"
   by (simp_all add: map_entry_def execute_simps)
 
@@ -295,20 +294,20 @@
   by (auto intro: success_intros simp add: map_entry_def)
 
 lemma crel_map_entryI [crel_intros]:
-  assumes "i < length a h" "h' = change a i (f (get_array a h ! i)) h" "r = a"
+  assumes "i < length a h" "h' = update a i (f (get_array a h ! i)) h" "r = a"
   shows "crel (map_entry i f a) h h' r"
   by (rule crelI) (insert assms, simp add: execute_simps)
 
 lemma crel_map_entryE [crel_elims]:
   assumes "crel (map_entry i f a) h h' r"
-  obtains "r = a" "h' = change a i (f (get_array a h ! i)) h" "i < length a h"
+  obtains "r = a" "h' = update a i (f (get_array a h ! i)) h" "i < length a h"
   using assms by (rule crelE)
     (erule successE, cases "i < length a h", simp_all add: execute_simps)
 
 lemma execute_swap [execute_simps]:
   "i < length a h \<Longrightarrow>
    execute (swap i x a) h =
-      Some (get_array a h ! i, change a i x h)"
+      Some (get_array a h ! i, update a i x h)"
   "i \<ge> length a h \<Longrightarrow> execute (swap i x a) h = None"
   by (simp_all add: swap_def execute_simps)
 
@@ -317,13 +316,13 @@
   by (auto intro: success_intros simp add: swap_def)
 
 lemma crel_swapI [crel_intros]:
-  assumes "i < length a h" "h' = change a i x h" "r = get_array a h ! i"
+  assumes "i < length a h" "h' = update a i x h" "r = get_array a h ! i"
   shows "crel (swap i x a) h h' r"
   by (rule crelI) (insert assms, simp add: execute_simps)
 
 lemma crel_swapE [crel_elims]:
   assumes "crel (swap i x a) h h' r"
-  obtains "r = get_array a h ! i" "h' = change a i x h" "i < length a h"
+  obtains "r = get_array a h ! i" "h' = update a i x h" "i < length a h"
   using assms by (rule crelE)
     (erule successE, cases "i < length a h", simp_all add: execute_simps)
 
@@ -357,7 +356,7 @@
   "of_list xs = make (List.length xs) (\<lambda>n. xs ! n)"
   by (rule Heap_eqI) (simp add: map_nth execute_simps)
 
-hide_const (open) new
+hide_const (open) update new of_list make len nth upd map_entry swap freeze
 
 
 subsection {* Code generator setup *}
@@ -407,24 +406,24 @@
   by (simp add: upd'_def upd_return)
 
 lemma [code]:
-  "map_entry i f a = do {
-     x \<leftarrow> nth a i;
-     upd i (f x) a
+  "Array.map_entry i f a = do {
+     x \<leftarrow> Array.nth a i;
+     Array.upd i (f x) a
    }"
   by (rule Heap_eqI) (simp add: bind_def guard_def map_entry_def execute_simps)
 
 lemma [code]:
-  "swap i x a = do {
-     y \<leftarrow> nth a i;
-     upd i x a;
+  "Array.swap i x a = do {
+     y \<leftarrow> Array.nth a i;
+     Array.upd i x a;
      return y
    }"
   by (rule Heap_eqI) (simp add: bind_def guard_def swap_def execute_simps)
 
 lemma [code]:
-  "freeze a = do {
-     n \<leftarrow> len a;
-     Heap_Monad.fold_map (\<lambda>i. nth a i) [0..<n]
+  "Array.freeze a = do {
+     n \<leftarrow> Array.len a;
+     Heap_Monad.fold_map (\<lambda>i. Array.nth a i) [0..<n]
    }"
 proof (rule Heap_eqI)
   fix h
@@ -441,12 +440,12 @@
     apply (simp add: length_def map_nth)
     done
   then have "execute (do {
-      n \<leftarrow> len a;
+      n \<leftarrow> Array.len a;
       Heap_Monad.fold_map (Array.nth a) [0..<n]
     }) h = Some (get_array a h, h)"
     by (auto intro: execute_bind_eq_SomeI simp add: execute_simps)
-  then show "execute (freeze a) h = execute (do {
-      n \<leftarrow> len a;
+  then show "execute (Array.freeze a) h = execute (do {
+      n \<leftarrow> Array.len a;
       Heap_Monad.fold_map (Array.nth a) [0..<n]
     }) h" by (simp add: execute_simps)
 qed
--- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Tue Jul 13 12:05:20 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy	Tue Jul 13 12:16:24 2010 +0200
@@ -14,17 +14,17 @@
 where
   "swap arr i j =
      do {
-       x \<leftarrow> nth arr i;
-       y \<leftarrow> nth arr j;
-       upd i y arr;
-       upd j x arr;
+       x \<leftarrow> Array.nth arr i;
+       y \<leftarrow> Array.nth arr j;
+       Array.upd i y arr;
+       Array.upd j x arr;
        return ()
      }"
 
 lemma crel_swapI [crel_intros]:
   assumes "i < Array.length a h" "j < Array.length a h"
     "x = get_array a h ! i" "y = get_array a h ! j"
-    "h' = Array.change a j x (Array.change a i y h)"
+    "h' = Array.update a j x (Array.update a i y h)"
   shows "crel (swap a i j) h h' r"
   unfolding swap_def using assms by (auto intro!: crel_intros)
 
@@ -41,7 +41,7 @@
   "part1 a left right p = (
      if (right \<le> left) then return right
      else do {
-       v \<leftarrow> nth a left;
+       v \<leftarrow> Array.nth a left;
        (if (v \<le> p) then (part1 a (left + 1) right p)
                     else (do { swap a left right;
   part1 a left (right - 1) p }))
@@ -228,9 +228,9 @@
 fun partition :: "nat array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat Heap"
 where
   "partition a left right = do {
-     pivot \<leftarrow> nth a right;
+     pivot \<leftarrow> Array.nth a right;
      middle \<leftarrow> part1 a left (right - 1) pivot;
-     v \<leftarrow> nth a middle;
+     v \<leftarrow> Array.nth a middle;
      m \<leftarrow> return (if (v \<le> pivot) then (middle + 1) else middle);
      swap a m right;
      return m
@@ -294,8 +294,8 @@
          else middle)"
     unfolding partition.simps
     by (elim crel_bindE crel_returnE crel_nthE crel_ifE crel_updE) simp
-  from swap have h'_def: "h' = Array.change a r (get_array a h1 ! rs)
-    (Array.change a rs (get_array a h1 ! r) h1)"
+  from swap have h'_def: "h' = Array.update a r (get_array a h1 ! rs)
+    (Array.update a rs (get_array a h1 ! r) h1)"
     unfolding swap_def
     by (elim crel_bindE crel_returnE crel_nthE crel_updE) simp
   from swap have in_bounds: "r < Array.length a h1 \<and> rs < Array.length a h1"
@@ -326,7 +326,7 @@
       with part_partitions[OF part] right_remains True
       have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
       with i_props h'_def in_bounds have "get_array a h' ! i \<le> get_array a h' ! rs"
-        unfolding Array.change_def Array.length_def by simp
+        unfolding Array.update_def Array.length_def by simp
     }
     moreover
     {
@@ -352,7 +352,7 @@
           by fastsimp
         with i_is True rs_equals right_remains h'_def
         show ?thesis using in_bounds
-          unfolding Array.change_def Array.length_def
+          unfolding Array.update_def Array.length_def
           by auto
       qed
     }
@@ -368,7 +368,7 @@
       from part_partitions[OF part] rs_equals right_remains i_is_left
       have "get_array a h1 ! i \<le> get_array a h' ! rs" by fastsimp
       with i_props h'_def have "get_array a h' ! i \<le> get_array a h' ! rs"
-        unfolding Array.change_def by simp
+        unfolding Array.update_def by simp
     }
     moreover
     {
@@ -388,7 +388,7 @@
         assume i_is: "i = r"
         from i_is False rs_equals right_remains h'_def
         show ?thesis using in_bounds
-          unfolding Array.change_def Array.length_def
+          unfolding Array.update_def Array.length_def
           by auto
       qed
     }
@@ -646,7 +646,7 @@
 subsection {* Example *}
 
 definition "qsort a = do {
-    k \<leftarrow> len a;
+    k \<leftarrow> Array.len a;
     quicksort a 0 (k - 1);
     return a
   }"
--- a/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy	Tue Jul 13 12:05:20 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy	Tue Jul 13 12:16:24 2010 +0200
@@ -8,14 +8,12 @@
 imports Imperative_HOL Subarray
 begin
 
-hide_const (open) swap rev
-
 fun swap :: "'a\<Colon>heap array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap" where
   "swap a i j = do {
-     x \<leftarrow> nth a i;
-     y \<leftarrow> nth a j;
-     upd i y a;
-     upd j x a;
+     x \<leftarrow> Array.nth a i;
+     y \<leftarrow> Array.nth a j;
+     Array.upd i y a;
+     Array.upd j x a;
      return ()
    }"
 
@@ -26,9 +24,6 @@
    }
    else return ())"
 
-notation (output) swap ("swap")
-notation (output) rev ("rev")
-
 declare swap.simps [simp del] rev.simps [simp del]
 
 lemma swap_pointwise: assumes "crel (swap a i j) h h' r"
--- a/src/HOL/Imperative_HOL/ex/SatChecker.thy	Tue Jul 13 12:05:20 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy	Tue Jul 13 12:16:24 2010 +0200
@@ -127,21 +127,21 @@
 unfolding array_ran_def Array.length_def by simp
 
 lemma array_ran_upd_array_Some:
-  assumes "cl \<in> array_ran a (Array.change a i (Some b) h)"
+  assumes "cl \<in> array_ran a (Array.update a i (Some b) h)"
   shows "cl \<in> array_ran a h \<or> cl = b"
 proof -
   have "set (get_array a h[i := Some b]) \<subseteq> insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert)
   with assms show ?thesis 
-    unfolding array_ran_def Array.change_def by fastsimp
+    unfolding array_ran_def Array.update_def by fastsimp
 qed
 
 lemma array_ran_upd_array_None:
-  assumes "cl \<in> array_ran a (Array.change a i None h)"
+  assumes "cl \<in> array_ran a (Array.update a i None h)"
   shows "cl \<in> array_ran a h"
 proof -
   have "set (get_array a h[i := None]) \<subseteq> insert None (set (get_array a h))" by (rule set_update_subset_insert)
   with assms show ?thesis
-    unfolding array_ran_def Array.change_def by auto
+    unfolding array_ran_def Array.update_def by auto
 qed
 
 definition correctArray :: "Clause list \<Rightarrow> Clause option array \<Rightarrow> heap \<Rightarrow> bool"
@@ -152,7 +152,7 @@
 lemma correctArray_update:
   assumes "correctArray rcs a h"
   assumes "correctClause rcs c" "sorted c" "distinct c"
-  shows "correctArray rcs a (Array.change a i (Some c) h)"
+  shows "correctArray rcs a (Array.update a i (Some c) h)"
   using assms
   unfolding correctArray_def
   by (auto dest:array_ran_upd_array_Some)
@@ -413,7 +413,7 @@
 definition get_clause :: "Clause option array \<Rightarrow> ClauseId \<Rightarrow> Clause Heap"
 where
   "get_clause a i = 
-       do { c \<leftarrow> nth a i;
+       do { c \<leftarrow> Array.nth a i;
            (case c of None \<Rightarrow> raise (''Clause not found'')
                     | Some x \<Rightarrow> return x)
        }"
@@ -440,11 +440,11 @@
   do {
      cli \<leftarrow> get_clause a i;
      result \<leftarrow> foldM (res_thm2 a) rs cli;
-     upd saveTo (Some result) a; 
+     Array.upd saveTo (Some result) a; 
      return rcs
   }"
-| "doProofStep2 a (Delete cid) rcs = do { upd cid None a; return rcs }"
-| "doProofStep2 a (Root cid clause) rcs = do { upd cid (Some (remdups (sort clause))) a; return (clause # rcs) }"
+| "doProofStep2 a (Delete cid) rcs = do { Array.upd cid None a; return rcs }"
+| "doProofStep2 a (Root cid clause) rcs = do { Array.upd cid (Some (remdups (sort clause))) a; return (clause # rcs) }"
 | "doProofStep2 a (Xstep cid1 cid2) rcs = raise ''MiniSatChecked.doProofStep: Xstep constructor found.''"
 | "doProofStep2 a (ProofDone b) rcs = raise ''MiniSatChecked.doProofStep: ProofDone constructor found.''"