1(*
2 * Copyright 2014, NICTA
3 *
4 * This software may be distributed and modified according to the terms of
5 * the BSD 2-Clause license. Note that NO WARRANTY is provided.
6 * See "LICENSE_BSD2.txt" for details.
7 *
8 * @TAG(NICTA_BSD)
9 *)
10
11theory LemmaBucket
12imports
13  HaskellLemmaBucket
14  SpecValid_R
15  SubMonadLib
16begin
17
18lemma corres_underlying_trivial:
19  "\<lbrakk> nf' \<Longrightarrow> no_fail P' f \<rbrakk> \<Longrightarrow> corres_underlying Id nf nf' (=) \<top> P' f f"
20  by (auto simp add: corres_underlying_def Id_def no_fail_def)
21
22lemma hoare_spec_gen_asm:
23  "\<lbrakk> F \<Longrightarrow> s \<turnstile> \<lbrace>P\<rbrace> f \<lbrace>Q\<rbrace> \<rbrakk> \<Longrightarrow> s \<turnstile> \<lbrace>P and K F\<rbrace> f \<lbrace>Q\<rbrace>"
24  "\<lbrakk> F \<Longrightarrow> s \<turnstile> \<lbrace>P\<rbrace> f' \<lbrace>Q\<rbrace>,\<lbrace>E\<rbrace> \<rbrakk> \<Longrightarrow> s \<turnstile> \<lbrace>P and K F\<rbrace> f' \<lbrace>Q\<rbrace>,\<lbrace>E\<rbrace>"
25  unfolding spec_valid_def spec_validE_def validE_def
26  apply (clarsimp simp only: pred_conj_def conj_assoc[symmetric]
27               intro!: hoare_gen_asm[unfolded pred_conj_def])+
28  done
29
30lemma spec_validE_fail:
31  "s \<turnstile> \<lbrace>P\<rbrace> fail \<lbrace>Q\<rbrace>,\<lbrace>E\<rbrace>"
32  by wp+
33
34lemma mresults_fail: "mresults fail = {}"
35  by (simp add: mresults_def fail_def)
36
37lemma gets_symb_exec_l:
38  "corres_underlying sr nf nf' dc P P' (gets f) (return x)"
39  by (simp add: corres_underlying_def return_def simpler_gets_def split_def)
40
41lemmas mapM_x_wp_inv = mapM_x_wp[where S=UNIV, simplified]
42
43lemma mapM_wp_inv:
44  "(\<And>x. \<lbrace>P\<rbrace> f x \<lbrace>\<lambda>rv. P\<rbrace>) \<Longrightarrow> \<lbrace>P\<rbrace> mapM f xs \<lbrace>\<lambda>rv. P\<rbrace>"
45  apply (rule  valid_return_unit)
46  apply (fold mapM_x_mapM)
47  apply (erule mapM_x_wp_inv)
48  done
49
50lemmas mapM_x_wp' = mapM_x_wp [OF _ subset_refl]
51
52lemma corres_underlying_similar:
53  "\<lbrakk> a = a'; b = b'; nf' \<Longrightarrow> no_fail \<top> (f a b) \<rbrakk>
54         \<Longrightarrow> corres_underlying Id nf nf' dc \<top> \<top> (f a b) (f a' b')"
55  by (simp add: corres_underlying_def no_fail_def, blast)
56
57lemma corres_underlying_gets_pre_lhs:
58  "(\<And>x. corres_underlying S nf nf' r (P x) P' (g x) g') \<Longrightarrow>
59  corres_underlying S nf nf' r (\<lambda>s. P (f s) s) P' (gets f >>= (\<lambda>x. g x)) g'"
60  apply (simp add: simpler_gets_def bind_def split_def corres_underlying_def)
61  apply force
62  done
63
64lemma mapM_x_inv_wp:
65  assumes x: "\<And>s. I s \<Longrightarrow> Q s"
66  assumes y: "\<And>x. x \<in> set xs \<Longrightarrow> \<lbrace>P\<rbrace> m x \<lbrace>\<lambda>rv. I\<rbrace>"
67  assumes z: "\<And>s. I s \<Longrightarrow> P s"
68  shows      "\<lbrace>I\<rbrace> mapM_x m xs \<lbrace>\<lambda>rv. Q\<rbrace>"
69  apply (rule hoare_post_imp)
70   apply (erule x)
71  apply (rule mapM_x_wp)
72   apply (rule hoare_pre_imp [OF _ y])
73    apply (erule z)
74   apply assumption
75  apply simp
76  done
77
78
79lemma mapM_x_accumulate_checks':
80  assumes P:  "\<And>x. x \<in> set xs' \<Longrightarrow> \<lbrace>\<top>\<rbrace> f x \<lbrace>\<lambda>rv. P x\<rbrace>"
81  assumes P': "\<And>x y. \<lbrakk> x \<in> set xs'; y \<in> set xs' \<rbrakk>
82                   \<Longrightarrow> \<lbrace>P y\<rbrace> f x \<lbrace>\<lambda>rv. P y\<rbrace>"
83  shows       "set xs \<subseteq> set xs' \<Longrightarrow>
84               \<lbrace>\<top>\<rbrace> mapM_x f xs \<lbrace>\<lambda>rv s. \<forall>x \<in> set xs. P x s\<rbrace>"
85  apply (induct xs)
86   apply (simp add: mapM_x_Nil)
87  apply (simp add: mapM_x_Cons)
88  apply (rule hoare_pre)
89   apply (wp mapM_x_wp'[OF P'])
90      apply blast
91     apply simp
92    apply assumption
93   apply simp
94   apply (rule P)
95   apply simp
96  apply simp
97  done
98
99lemmas mapM_x_accumulate_checks
100    = mapM_x_accumulate_checks'[OF _ _ subset_refl]
101
102(* Other *)
103
104lemma isRight_rel_sum_comb2:
105  "\<lbrakk> (f \<oplus> r) v v'; isRight v' \<rbrakk>
106       \<Longrightarrow> isRight v \<and> r (theRight v) (theRight v')"
107  by (clarsimp simp: isRight_def)
108
109lemma isRight_case_sum: "isRight x \<Longrightarrow> case_sum f g x = g (theRight x)"
110  by (clarsimp simp add: isRight_def)
111
112lemma enumerate_append:"enumerate i (xs @ ys) = enumerate i xs @ enumerate (i + length xs) ys"
113  apply (induct xs arbitrary:ys i)
114   apply clarsimp+
115  done
116
117lemma enumerate_bound:"(a, b) \<in> set (enumerate n xs) \<Longrightarrow> a < n + length xs"
118  by (metis add.commute in_set_enumerate_eq prod.sel(1))
119
120lemma enumerate_exceed:"(n + length xs, b) \<notin> set (enumerate n xs)"
121  by (metis enumerate_bound less_not_refl)
122
123lemma all_pair_unwrap:"(\<forall>a. P (fst a) (snd a)) = (\<forall>a b. P a b)"
124  by force
125
126lemma if_fold[simp]:"(if P then Q else if P then R else S) = (if P then Q else S)"
127  by presburger
128
129lemma disjoint_subset_both:"\<lbrakk>A' \<subseteq> A; B' \<subseteq> B; A \<inter> B = {}\<rbrakk> \<Longrightarrow> A' \<inter> B' = {}"
130  by blast
131
132lemma union_split: "\<lbrakk>A \<inter> C = {}; B \<inter> C = {}\<rbrakk> \<Longrightarrow> (A \<union> B) \<inter> C = {}"
133  by (simp add: inf_sup_distrib2)
134
135lemma dom_expand: "dom (\<lambda>x. if P x then Some y else None) = {x. P x}"
136  using if_option_Some by fastforce
137
138lemma range_translate: "(range f = range g) = ((\<forall>x. \<exists>y. f x = g y) \<and> (\<forall>x. \<exists>y. f y = g x))"
139  by (rule iffI,
140       rule conjI,
141        clarsimp,
142        blast,
143       clarsimp,
144       metis f_inv_into_f range_eqI,
145      clarsimp,
146      subst set_eq_subset,
147      rule conjI,
148       clarsimp,
149       rename_tac arg,
150       erule_tac x=arg and P="\<lambda>x. (\<exists>y. f x = g y)" in allE,
151       clarsimp,
152      clarsimp,
153      rename_tac arg,
154      erule_tac x=arg and P="\<lambda>x. (\<exists>y. f y = g x)" in allE,
155      clarsimp,
156      metis range_eqI)
157
158lemma ran_expand: "\<exists>x. P x \<Longrightarrow> ran (\<lambda>x. if P x then Some y else None) = {y}"
159  by (rule subset_antisym,
160       (clarsimp simp:ran_def)+)
161
162lemma map_upd_expand: "f(x \<mapsto> y) = f ++ (\<lambda>z. if z = x then Some y else None)"
163  by (rule ext, rename_tac w,
164      case_tac "w = x",
165       simp,
166      simp add:map_add_def)
167
168lemma map_upd_subI: "\<lbrakk>f \<subseteq>\<^sub>m g; f x = None\<rbrakk> \<Longrightarrow> f \<subseteq>\<^sub>m g(x \<mapsto> y)"
169  by (rule_tac f="\<lambda>i. if i = x then Some y else None" in map_add_le_mapE,
170      simp add:map_le_def,
171      rule ballI, rename_tac a,
172      rule conjI,
173       erule_tac x=x in ballE,
174        clarsimp,
175        erule disjE,
176         clarsimp,
177        clarsimp simp:map_add_def,
178       clarsimp,
179       erule disjE,
180        clarsimp,
181       clarsimp simp:map_add_def,
182      clarsimp simp:map_add_def,
183      erule_tac x=a in ballE,
184       erule disjE,
185        (case_tac "g a"; simp_all),
186       clarsimp+)
187
188lemma all_ext: "\<forall>x. f x = g x \<Longrightarrow> f = g"
189  by presburger
190
191lemma conjI2: "\<lbrakk>B; B \<longrightarrow> A\<rbrakk> \<Longrightarrow> A \<and> B"
192  by auto
193
194(* Trivial lemmas for dealing with messy CNode obligations. *)
195lemma Least2: "\<lbrakk>\<not>P 0; \<not>P 1; P (2::nat)\<rbrakk> \<Longrightarrow> Least P = 2"
196  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
197lemma Least3: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; P (3::nat)\<rbrakk> \<Longrightarrow> Least P = 3"
198  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
199lemma Least4: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; P (4::nat)\<rbrakk> \<Longrightarrow> Least P = 4"
200  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
201lemma Least5: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; P (5::nat)\<rbrakk> \<Longrightarrow> Least P = 5"
202  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
203lemma Least6: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; P (6::nat)\<rbrakk> \<Longrightarrow> Least P = 6"
204  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
205lemma Least7: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; P (7::nat)\<rbrakk> \<Longrightarrow> Least P = 7"
206  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
207lemma Least8: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; P (8::nat)\<rbrakk> \<Longrightarrow> Least P = 8"
208  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
209lemma Least9: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; P (9::nat)\<rbrakk> \<Longrightarrow> Least P = 9"
210  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
211lemma Least10: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; P (10::nat)\<rbrakk> \<Longrightarrow> Least P
212                 = 10"
213  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
214lemma Least11: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; P (11::nat)\<rbrakk> \<Longrightarrow>
215                 Least P = 11"
216  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
217lemma Least12: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; P
218                 (12::nat)\<rbrakk> \<Longrightarrow> Least P = 12"
219  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
220lemma Least13: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; P
221                 (13::nat)\<rbrakk> \<Longrightarrow> Least P = 13"
222  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
223lemma Least14: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
224                 13; P (14::nat)\<rbrakk> \<Longrightarrow> Least P = 14"
225  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
226lemma Least15: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
227                 13; \<not>P 14; P (15::nat)\<rbrakk> \<Longrightarrow> Least P = 15"
228  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
229lemma Least16: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
230                 13; \<not>P 14; \<not>P 15; P (16::nat)\<rbrakk> \<Longrightarrow> Least P = 16"
231  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
232lemma Least17: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
233                 13; \<not>P 14; \<not>P 15; \<not>P 16; P (17::nat)\<rbrakk> \<Longrightarrow> Least P = 17"
234  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
235lemma Least18: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
236                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; P (18::nat)\<rbrakk> \<Longrightarrow> Least P = 18"
237  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
238lemma Least19: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
239                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; P (19::nat)\<rbrakk> \<Longrightarrow> Least P = 19"
240  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
241lemma Least20: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
242                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; P (20::nat)\<rbrakk> \<Longrightarrow> Least P = 20"
243  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
244lemma Least21: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
245                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; P (21::nat)\<rbrakk> \<Longrightarrow> Least P = 21"
246  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
247lemma Least22: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
248                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; \<not>P 21; P (22::nat)\<rbrakk> \<Longrightarrow> Least P
249                 = 22"
250  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
251lemma Least23: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
252                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; \<not>P 21; \<not>P 22; P (23::nat)\<rbrakk> \<Longrightarrow>
253                 Least P = 23"
254  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
255lemma Least24: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
256                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; \<not>P 21; \<not>P 22; \<not>P 23; P
257                 (24::nat)\<rbrakk> \<Longrightarrow> Least P = 24"
258  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
259lemma Least25: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
260                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; \<not>P 21; \<not>P 22; \<not>P 23; \<not>P 24; P
261                 (25::nat)\<rbrakk> \<Longrightarrow> Least P = 25"
262  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
263lemma Least26: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
264                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; \<not>P 21; \<not>P 22; \<not>P 23; \<not>P 24; \<not>P
265                 25; P (26::nat)\<rbrakk> \<Longrightarrow> Least P = 26"
266  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
267lemma Least27: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
268                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; \<not>P 21; \<not>P 22; \<not>P 23; \<not>P 24; \<not>P
269                 25; \<not>P 26; P (27::nat)\<rbrakk> \<Longrightarrow> Least P = 27"
270  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
271lemma Least28: "\<lbrakk>\<not>P 0; \<not>P 1; \<not>P 2; \<not>P 3; \<not>P 4; \<not>P 5; \<not>P 6; \<not>P 7; \<not>P 8; \<not>P 9; \<not>P 10; \<not>P 11; \<not>P 12; \<not>P
272                 13; \<not>P 14; \<not>P 15; \<not>P 16; \<not>P 17; \<not>P 18; \<not>P 19; \<not>P 20; \<not>P 21; \<not>P 22; \<not>P 23; \<not>P 24; \<not>P
273                 25; \<not>P 26; \<not>P 27; P (28::nat)\<rbrakk> \<Longrightarrow> Least P = 28"
274  by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3))
275
276lemma map_add_discard: "\<not> cond x \<Longrightarrow> (f ++ (\<lambda>x. if cond x then (g x) else None)) x = f x"
277  by (simp add: map_add_def)
278
279lemma dom_split:"\<lbrakk>\<forall>x \<in> S. \<exists>y. f x = Some y; \<forall>x. x \<notin> S \<longrightarrow> f x = None\<rbrakk> \<Longrightarrow> dom f = S"
280  by (auto simp:dom_def)
281
282lemma map_set_in: "x \<in> f ` S = (\<exists>y\<in>S. f y = x)"
283  by blast
284
285lemma map_length_split:
286  "map (length \<circ> (\<lambda>(a, b). P a b # map (f a b) (Q a b))) xs = map (\<lambda>(a, b). 1 + length (Q a b)) xs"
287  by clarsimp
288
289lemma sum_suc: "(\<Sum>x \<leftarrow> xs. Suc (f x)) = length xs + (\<Sum>x \<leftarrow> xs. f x)"
290  apply (induct xs)
291   by clarsimp+
292
293lemma sum_suc_pair: "(\<Sum>(a, b) \<leftarrow> xs. Suc (f a b)) = length xs + (\<Sum>(a, b) \<leftarrow> xs. f a b)"
294  apply (induct xs)
295   by clarsimp+
296
297lemma fold_add_sum: "fold (+) ((map (\<lambda>(a, b). f a b) xs)::nat list) 0 = (\<Sum>(a, b) \<leftarrow> xs. f a b)"
298  apply (subst fold_plus_sum_list_rev)
299  apply (subst sum_list_rev)
300  by clarsimp
301
302lemma set_of_enumerate:"card (set (enumerate n xs)) = length xs"
303  by (metis distinct_card distinct_enumerate length_enumerate)
304
305lemma collapse_fst: "fst ` (\<lambda>x. (f x, g x)) ` s = f ` s"
306  by force
307
308lemma collapse_fst2: "fst ` (\<lambda>(x, y). (f x, g y)) ` s = (\<lambda>x. f (fst x)) ` s"
309  by force
310
311lemma collapse_fst3: "(\<lambda>x. f (fst x)) ` set (enumerate n xs) = f ` set [n..<n + length xs]"
312  by (metis image_image list.set_map map_fst_enumerate)
313
314lemma card_of_dom_bounded:
315  fixes f :: "'a \<Rightarrow> 'b option"
316  assumes "finite (UNIV::'a set)"
317  shows "card (dom f) \<le> CARD('a)"
318  by (simp add: assms card_mono)
319
320lemma third_in: "(a, b, c) \<in> S \<Longrightarrow> c \<in> (snd \<circ> snd) ` S"
321  by (metis (erased, hide_lams) map_set_in image_comp snd_conv)
322
323lemma third_in2: "(a \<in> (snd \<circ> snd) ` (set (enumerate i xs))) = (a \<in> snd ` (set xs))"
324  by (metis map_map map_snd_enumerate set_map)
325
326lemma map_of_enum: "map_of (enumerate n xs) x = Some y \<Longrightarrow> y \<in> set xs"
327  apply (clarsimp)
328  by (metis enumerate_eq_zip in_set_zipE)
329
330lemma map_of_append:
331  "(map_of xs ++ map_of ys) x = (case map_of ys x of None \<Rightarrow> map_of xs x | Some x' \<Rightarrow> Some x')"
332  by (simp add: map_add_def)
333
334lemma map_of_append2:
335  "(map_of xs ++ map_of ys ++ map_of zs) x =
336     (case map_of zs x of None \<Rightarrow> (case map_of ys x of None \<Rightarrow> map_of xs x
337                                                     | Some x' \<Rightarrow> Some x')
338                        | Some x' \<Rightarrow> Some x')"
339  by (simp add: map_add_def)
340
341lemma map_of_in_set_map: "map_of (map (\<lambda>(n, y). (f n, y)) xs) x = Some z \<Longrightarrow> z \<in> snd ` set xs"
342  proof -
343    assume "map_of (map (\<lambda>(n, y). (f n, y)) xs) x = Some z"
344    hence "(x, z) \<in> (\<lambda>(uu, y). (f uu, y)) ` set xs" using map_of_SomeD by fastforce
345    thus "z \<in> snd ` set xs" using map_set_in by fastforce
346  qed
347
348lemma pair_in_enum: "(a, b) \<in> set (enumerate x ys) \<Longrightarrow> b \<in> set ys"
349  by (metis enumerate_eq_zip in_set_zip2)
350
351lemma distinct_inj:
352  "inj f \<Longrightarrow> distinct xs = distinct (map f xs)"
353  apply (induct xs)
354   apply simp
355  apply (simp add: inj_image_mem_iff)
356  done
357
358lemma distinct_map_via_ran: "distinct (map fst xs) \<Longrightarrow> ran (map_of xs) = set (map snd xs)"
359  apply (cut_tac xs="map fst xs" and ys="map snd xs" in ran_map_of_zip[symmetric])
360    apply clarsimp+
361  by (simp add: ran_distinct)
362
363lemma in_ran_in_set: "x \<in> ran (map_of xs) \<Longrightarrow> x \<in> set (map snd xs)"
364  by (metis (mono_tags, hide_lams) map_set_in map_of_SomeD ranE set_map snd_conv)
365
366lemma in_ran_map_app: "x \<in> ran (xs ++ ys ++ zs) \<Longrightarrow> x \<in> ran xs \<or> x \<in> ran ys \<or> x \<in> ran zs"
367  proof -
368    assume a1: "x \<in> ran (xs ++ ys ++ zs)"
369    obtain bb :: "'a \<Rightarrow> ('b \<Rightarrow> 'a option) \<Rightarrow> 'b" where
370      "\<forall>x0 x1. (\<exists>v2. x1 v2 = Some x0) = (x1 (bb x0 x1) = Some x0)"
371      by moura
372    hence f2: "\<forall>f a. (\<not> (\<exists>b. f b = Some a) \<or> f (bb a f) = Some a) \<and> ((\<exists>b. f b = Some a) \<or> (\<forall>b. f b \<noteq> Some a))"
373      by blast
374    have "\<exists>b. (xs ++ ys ++ zs) b = Some x"
375      using a1 by (simp add: ran_def)
376    hence f3: "(xs ++ ys ++ zs) (bb x (xs ++ ys ++ zs)) = Some x"
377      using f2 by meson
378    { assume "ys (bb x (xs ++ ys ++ zs)) \<noteq> None \<or> xs (bb x (xs ++ ys ++ zs)) \<noteq> Some x"
379      { assume "ys (bb x (xs ++ ys ++ zs)) \<noteq> Some x \<and> (ys (bb x (xs ++ ys ++ zs)) \<noteq> None \<or> xs (bb x (xs ++ ys ++ zs)) \<noteq> Some x)"
380        hence "\<exists>b. zs b = Some x"
381          using f3 by auto
382        hence ?thesis
383          by (simp add: ran_def) }
384      hence ?thesis
385        using ran_def by fastforce }
386    thus ?thesis
387      using ran_def by fastforce
388  qed
389
390lemma none_some_map: "None \<notin> S \<Longrightarrow> Some x \<in> S = (x \<in> the ` S)"
391  apply (rule iffI)
392   apply force
393  apply (subst in_these_eq[symmetric])
394  apply (clarsimp simp:Option.these_def)
395  apply (case_tac "\<exists>y. xa = Some y")
396   by clarsimp+
397
398lemma none_some_map2: "the ` Set.filter (\<lambda>s. \<not> Option.is_none s) (range f) = ran f"
399  apply (rule subset_antisym)
400   apply clarsimp
401   apply (case_tac "f x", simp_all)
402   apply (simp add: ranI)
403  apply clarsimp
404  apply (subst none_some_map[symmetric])
405   apply clarsimp+
406  apply (erule ranE)
407  by (metis range_eqI)
408
409lemma prop_map_of_prop:"\<lbrakk>\<forall>z \<in> set xs. P (g z); map_of (map (\<lambda>x. (f x, g x)) xs) y = Some a\<rbrakk> \<Longrightarrow> P a"
410  using map_of_SomeD by fastforce
411
412lemma range_subsetI2: "\<forall>y\<in>A. \<exists>x. f x = y \<Longrightarrow> A \<subseteq> range f"
413 by fast
414
415lemma insert_strip: "x \<noteq> y \<Longrightarrow> (x \<in> insert y S) = (x \<in> S)"
416  by simp
417
418lemma dom_map_add: "dom ys = A \<Longrightarrow> dom (xs ++ ys) = A \<union> dom xs"
419  by simp
420
421lemma set_compre_unwrap: "({x. P x} \<subseteq> S) = (\<forall>x. P x \<longrightarrow> x \<in> S)"
422  by blast
423
424lemma map_add_same: "\<lbrakk>xs = ys; zs = ws\<rbrakk> \<Longrightarrow> xs ++ zs = ys ++ ws"
425  by simp
426
427lemma map_add_find_left: "n k = None \<Longrightarrow> (m ++ n) k = m k"
428  by (simp add:map_add_def)
429
430lemma map_length_split_triple:
431  "map (length \<circ> (\<lambda>(a, b, c). P a b c # map (f a b c) (Q a b c))) xs =
432     map (\<lambda>(a, b, c). 1 + length (Q a b c)) xs"
433  by fastforce
434
435lemma sum_suc_triple: "(\<Sum>(a, b, c)\<leftarrow>xs. Suc (f a b c)) = length xs + (\<Sum>(a, b, c)\<leftarrow>xs. f a b c)"
436  by (induct xs; clarsimp)
437
438lemma sum_enumerate: "(\<Sum>(a, b)\<leftarrow>enumerate n xs. P b) = (\<Sum>b\<leftarrow>xs. P b)"
439  by (induct xs arbitrary:n; clarsimp)
440
441lemma dom_map_fold:"dom (fold (++) (map (\<lambda>x. [f x \<mapsto> g x]) xs) ms) = dom ms \<union> set (map f xs)"
442  by (induct xs arbitrary:f g ms; clarsimp)
443
444lemma list_ran_prop:"map_of (map (\<lambda>x. (f x, g x)) xs) i = Some t \<Longrightarrow> \<exists>x \<in> set xs. g x = t"
445  by (induct xs arbitrary:f g t i; clarsimp split:if_split_asm)
446
447lemma in_set_enumerate_eq2:"(a, b) \<in> set (enumerate n xs) \<Longrightarrow> (b = xs ! (a - n))"
448  by (simp add: in_set_enumerate_eq)
449
450lemma subset_eq_notI: "\<lbrakk>a\<in> B;a\<notin> C\<rbrakk> \<Longrightarrow> \<not> B \<subseteq> C"
451  by auto
452
453lemma nat_divide_less_eq:
454  fixes b :: nat
455  shows "0 < c \<Longrightarrow> (b div c < a) = (b < a * c)"
456  using td_gal_lt by blast
457
458lemma strengthen_imp_same_first_conj:
459  "(b \<and> (a \<longrightarrow> c) \<and> (a' \<longrightarrow> c')) \<Longrightarrow> ((a \<longrightarrow> b \<and> c) \<and> (a' \<longrightarrow> b \<and> c'))"
460  by blast
461
462lemma conj_impD:
463  "a \<and> b \<Longrightarrow> a \<longrightarrow> b"
464  by blast
465
466lemma set_list_mem_nonempty:
467  "x \<in> set xs \<Longrightarrow> xs \<noteq> []"
468  by auto
469
470lemma strenghten_False_imp:
471  "\<not>P \<Longrightarrow> P \<longrightarrow> Q"
472  by blast
473
474lemma foldl_fun_or_alt:
475  "foldl (\<lambda>x y. x \<or> f y) b ls = foldl (\<or>) b (map f ls)"
476  apply (induct ls)
477   apply clarsimp
478  apply clarsimp
479  by (simp add: foldl_map)
480
481end
482