src/HOL/Bali/Basis.thy
author wenzelm
Thu, 28 Feb 2013 13:24:51 +0100
changeset 52441 0e71a248cacb
parent 46022 2dd44cd8f963
child 52854 9e7d1c139569
permissions -rw-r--r--
marginalized historic strip_tac;
wenzelm@12857
     1
(*  Title:      HOL/Bali/Basis.thy
schirmer@12854
     2
    Author:     David von Oheimb
schirmer@12854
     3
*)
schirmer@12854
     4
header {* Definitions extending HOL as logical basis of Bali *}
schirmer@12854
     5
wenzelm@46022
     6
theory Basis
wenzelm@46022
     7
imports Main "~~/src/HOL/Library/Old_Recdef"
wenzelm@46022
     8
begin
schirmer@12854
     9
schirmer@12854
    10
section "misc"
schirmer@12854
    11
wenzelm@52441
    12
ML {* fun strip_tac i = REPEAT (resolve_tac [impI, allI] i) *}
wenzelm@52441
    13
schirmer@12854
    14
declare split_if_asm  [split] option.split [split] option.split_asm [split]
wenzelm@24019
    15
declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
schirmer@12854
    16
declare if_weak_cong [cong del] option.weak_case_cong [cong del]
paulson@18447
    17
declare length_Suc_conv [iff]
paulson@18447
    18
schirmer@12854
    19
lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}"
wenzelm@46022
    20
  by auto
schirmer@12854
    21
wenzelm@46022
    22
lemma subset_insertD: "A \<subseteq> insert x B \<Longrightarrow> A \<subseteq> B \<and> x \<notin> A \<or> (\<exists>B'. A = insert x B' \<and> B' \<subseteq> B)"
wenzelm@46022
    23
  apply (case_tac "x \<in> A")
wenzelm@46022
    24
   apply (rule disjI2)
wenzelm@46022
    25
   apply (rule_tac x = "A - {x}" in exI)
wenzelm@46022
    26
   apply fast+
wenzelm@46022
    27
  done
schirmer@12854
    28
wenzelm@46022
    29
abbreviation nat3 :: nat  ("3") where "3 \<equiv> Suc 2"
wenzelm@46022
    30
abbreviation nat4 :: nat  ("4") where "4 \<equiv> Suc 3"
schirmer@12854
    31
nipkow@13867
    32
(* irrefl_tranclI in Transitive_Closure.thy is more general *)
wenzelm@46022
    33
lemma irrefl_tranclI': "r\<inverse> \<inter> r\<^sup>+ = {} \<Longrightarrow> \<forall>x. (x, x) \<notin> r\<^sup>+"
wenzelm@46022
    34
  by (blast elim: tranclE dest: trancl_into_rtrancl)
nipkow@13867
    35
schirmer@12854
    36
wenzelm@46022
    37
lemma trancl_rtrancl_trancl: "\<lbrakk>(x, y) \<in> r\<^sup>+; (y, z) \<in> r\<^sup>*\<rbrakk> \<Longrightarrow> (x, z) \<in> r\<^sup>+"
wenzelm@46022
    38
  by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2)
schirmer@12854
    39
wenzelm@46022
    40
lemma rtrancl_into_trancl3: "\<lbrakk>(a, b) \<in> r\<^sup>*; a \<noteq> b\<rbrakk> \<Longrightarrow> (a, b) \<in> r\<^sup>+"
wenzelm@46022
    41
  apply (drule rtranclD)
wenzelm@46022
    42
  apply auto
wenzelm@46022
    43
  done
schirmer@12854
    44
wenzelm@46022
    45
lemma rtrancl_into_rtrancl2: "\<lbrakk>(a, b) \<in>  r; (b, c) \<in> r\<^sup>*\<rbrakk> \<Longrightarrow> (a, c) \<in> r\<^sup>*"
wenzelm@46022
    46
  by (auto intro: rtrancl_trans)
schirmer@12854
    47
schirmer@12854
    48
lemma triangle_lemma:
wenzelm@46022
    49
  assumes unique: "\<And>a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b = c"
wenzelm@46022
    50
    and ax: "(a,x)\<in>r\<^sup>*" and ay: "(a,y)\<in>r\<^sup>*"
wenzelm@46022
    51
  shows "(x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"
wenzelm@46022
    52
  using ax ay
wenzelm@46022
    53
proof (induct rule: converse_rtrancl_induct)
wenzelm@46022
    54
  assume "(x,y)\<in>r\<^sup>*"
wenzelm@46022
    55
  then show ?thesis by blast
wenzelm@46022
    56
next
wenzelm@46022
    57
  fix a v
wenzelm@46022
    58
  assume a_v_r: "(a, v) \<in> r"
wenzelm@46022
    59
    and v_x_rt: "(v, x) \<in> r\<^sup>*"
wenzelm@46022
    60
    and a_y_rt: "(a, y) \<in> r\<^sup>*"
wenzelm@46022
    61
    and hyp: "(v, y) \<in> r\<^sup>* \<Longrightarrow> (x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
wenzelm@46022
    62
  from a_y_rt show "(x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
wenzelm@46022
    63
  proof (cases rule: converse_rtranclE)
wenzelm@46022
    64
    assume "a = y"
wenzelm@46022
    65
    with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*"
wenzelm@46022
    66
      by (auto intro: rtrancl_trans)
wenzelm@46022
    67
    then show ?thesis by blast
schirmer@12854
    68
  next
wenzelm@46022
    69
    fix w
wenzelm@46022
    70
    assume a_w_r: "(a, w) \<in> r"
wenzelm@46022
    71
      and w_y_rt: "(w, y) \<in> r\<^sup>*"
wenzelm@46022
    72
    from a_v_r a_w_r unique have "v=w" by auto
wenzelm@46022
    73
    with w_y_rt hyp show ?thesis by blast
schirmer@12854
    74
  qed
schirmer@12854
    75
qed
schirmer@12854
    76
schirmer@12854
    77
wenzelm@46022
    78
lemma rtrancl_cases:
wenzelm@46022
    79
  assumes "(a,b)\<in>r\<^sup>*"
wenzelm@46022
    80
  obtains (Refl) "a = b"
wenzelm@46022
    81
    | (Trancl) "(a,b)\<in>r\<^sup>+"
wenzelm@46022
    82
  apply (rule rtranclE [OF assms])
wenzelm@46022
    83
   apply (auto dest: rtrancl_into_trancl1)
wenzelm@46022
    84
  done
schirmer@12854
    85
wenzelm@46022
    86
lemma Ball_weaken: "\<lbrakk>Ball s P; \<And> x. P x\<longrightarrow>Q x\<rbrakk>\<Longrightarrow>Ball s Q"
wenzelm@46022
    87
  by auto
schirmer@12854
    88
wenzelm@46022
    89
lemma finite_SetCompr2:
wenzelm@46022
    90
  "finite (Collect P) \<Longrightarrow> \<forall>y. P y \<longrightarrow> finite (range (f y)) \<Longrightarrow>
wenzelm@46022
    91
    finite {f y x |x y. P y}"
wenzelm@46022
    92
  apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (\<lambda>y. range (f y))")
wenzelm@46022
    93
   prefer 2 apply fast
wenzelm@46022
    94
  apply (erule ssubst)
wenzelm@46022
    95
  apply (erule finite_UN_I)
wenzelm@46022
    96
  apply fast
wenzelm@46022
    97
  done
schirmer@12854
    98
wenzelm@46022
    99
lemma list_all2_trans: "\<forall>a b c. P1 a b \<longrightarrow> P2 b c \<longrightarrow> P3 a c \<Longrightarrow>
wenzelm@46022
   100
    \<forall>xs2 xs3. list_all2 P1 xs1 xs2 \<longrightarrow> list_all2 P2 xs2 xs3 \<longrightarrow> list_all2 P3 xs1 xs3"
wenzelm@46022
   101
  apply (induct_tac xs1)
wenzelm@46022
   102
   apply simp
wenzelm@46022
   103
  apply (rule allI)
wenzelm@46022
   104
  apply (induct_tac xs2)
wenzelm@46022
   105
   apply simp
wenzelm@46022
   106
  apply (rule allI)
wenzelm@46022
   107
  apply (induct_tac xs3)
wenzelm@46022
   108
   apply auto
wenzelm@46022
   109
  done
schirmer@12854
   110
schirmer@12854
   111
schirmer@12854
   112
section "pairs"
schirmer@12854
   113
wenzelm@46022
   114
lemma surjective_pairing5:
wenzelm@46022
   115
  "p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))),
wenzelm@46022
   116
    snd (snd (snd (snd p))))"
wenzelm@46022
   117
  by auto
schirmer@12854
   118
wenzelm@46022
   119
lemma fst_splitE [elim!]:
wenzelm@46022
   120
  assumes "fst s' = x'"
wenzelm@46022
   121
  obtains x s where "s' = (x,s)" and "x = x'"
wenzelm@46022
   122
  using assms by (cases s') auto
schirmer@12854
   123
wenzelm@46022
   124
lemma fst_in_set_lemma: "(x, y) : set l \<Longrightarrow> x : fst ` set l"
wenzelm@46022
   125
  by (induct l) auto
schirmer@12854
   126
schirmer@12854
   127
schirmer@12854
   128
section "quantifiers"
schirmer@12854
   129
wenzelm@46022
   130
lemma All_Ex_refl_eq2 [simp]: "(\<forall>x. (\<exists>b. x = f b \<and> Q b) \<longrightarrow> P x) = (\<forall>b. Q b \<longrightarrow> P (f b))"
wenzelm@46022
   131
  by auto
schirmer@12854
   132
wenzelm@46022
   133
lemma ex_ex_miniscope1 [simp]: "(\<exists>w v. P w v \<and> Q v) = (\<exists>v. (\<exists>w. P w v) \<and> Q v)"
wenzelm@46022
   134
  by auto
schirmer@12854
   135
wenzelm@46022
   136
lemma ex_miniscope2 [simp]: "(\<exists>v. P v \<and> Q \<and> R v) = (Q \<and> (\<exists>v. P v \<and> R v))"
wenzelm@46022
   137
  by auto
schirmer@12854
   138
schirmer@12854
   139
lemma ex_reorder31: "(\<exists>z x y. P x y z) = (\<exists>x y z. P x y z)"
wenzelm@46022
   140
  by auto
schirmer@12854
   141
wenzelm@46022
   142
lemma All_Ex_refl_eq1 [simp]: "(\<forall>x. (\<exists>b. x = f b) \<longrightarrow> P x) = (\<forall>b. P (f b))"
wenzelm@46022
   143
  by auto
schirmer@12854
   144
schirmer@12854
   145
schirmer@12854
   146
section "sums"
schirmer@12854
   147
wenzelm@36176
   148
hide_const In0 In1
schirmer@12854
   149
wenzelm@35078
   150
notation sum_case  (infixr "'(+')"80)
schirmer@12854
   151
wenzelm@46022
   152
primrec the_Inl :: "'a + 'b \<Rightarrow> 'a"
wenzelm@38219
   153
  where "the_Inl (Inl a) = a"
wenzelm@38219
   154
wenzelm@46022
   155
primrec the_Inr :: "'a + 'b \<Rightarrow> 'b"
wenzelm@38219
   156
  where "the_Inr (Inr b) = b"
schirmer@12854
   157
schirmer@12854
   158
datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c
schirmer@12854
   159
wenzelm@46022
   160
primrec the_In1 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'a"
wenzelm@38219
   161
  where "the_In1 (In1 a) = a"
wenzelm@38219
   162
wenzelm@46022
   163
primrec the_In2 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'b"
wenzelm@38219
   164
  where "the_In2 (In2 b) = b"
wenzelm@38219
   165
wenzelm@46022
   166
primrec the_In3 :: "('a, 'b, 'c) sum3 \<Rightarrow> 'c"
wenzelm@38219
   167
  where "the_In3 (In3 c) = c"
schirmer@12854
   168
wenzelm@46022
   169
abbreviation In1l :: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
wenzelm@46022
   170
  where "In1l e \<equiv> In1 (Inl e)"
schirmer@12854
   171
wenzelm@46022
   172
abbreviation In1r :: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
wenzelm@46022
   173
  where "In1r c \<equiv> In1 (Inr c)"
wenzelm@35078
   174
wenzelm@35078
   175
abbreviation the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al"
wenzelm@46022
   176
  where "the_In1l \<equiv> the_Inl \<circ> the_In1"
wenzelm@35078
   177
wenzelm@35078
   178
abbreviation the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar"
wenzelm@46022
   179
  where "the_In1r \<equiv> the_Inr \<circ> the_In1"
schirmer@13688
   180
schirmer@12854
   181
ML {*
wenzelm@27226
   182
fun sum3_instantiate ctxt thm = map (fn s =>
wenzelm@46022
   183
  simplify (simpset_of ctxt delsimps [@{thm not_None_eq}])
wenzelm@27239
   184
    (read_instantiate ctxt [(("t", 0), "In" ^ s ^ " ?x")] thm)) ["1l","2","3","1r"]
schirmer@12854
   185
*}
schirmer@12854
   186
(* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *)
schirmer@12854
   187
schirmer@12854
   188
schirmer@12854
   189
section "quantifiers for option type"
schirmer@12854
   190
schirmer@12854
   191
syntax
wenzelm@46022
   192
  "_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool"   ("(3! _:_:/ _)" [0,0,10] 10)
wenzelm@46022
   193
  "_Oex"  :: "[pttrn, 'a option, bool] \<Rightarrow> bool"   ("(3? _:_:/ _)" [0,0,10] 10)
schirmer@12854
   194
schirmer@12854
   195
syntax (symbols)
wenzelm@46022
   196
  "_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
wenzelm@46022
   197
  "_Oex"  :: "[pttrn, 'a option, bool] \<Rightarrow> bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
schirmer@12854
   198
schirmer@12854
   199
translations
wenzelm@46022
   200
  "\<forall>x\<in>A: P" \<rightleftharpoons> "\<forall>x\<in>CONST Option.set A. P"
wenzelm@46022
   201
  "\<exists>x\<in>A: P" \<rightleftharpoons> "\<exists>x\<in>CONST Option.set A. P"
wenzelm@46022
   202
schirmer@12854
   203
nipkow@19323
   204
section "Special map update"
nipkow@19323
   205
nipkow@19323
   206
text{* Deemed too special for theory Map. *}
nipkow@19323
   207
wenzelm@46022
   208
definition chg_map :: "('b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> ('a \<rightharpoonup> 'b) \<Rightarrow> ('a \<rightharpoonup> 'b)"
wenzelm@46022
   209
  where "chg_map f a m = (case m a of None \<Rightarrow> m | Some b \<Rightarrow> m(a\<mapsto>f b))"
nipkow@19323
   210
wenzelm@46022
   211
lemma chg_map_new[simp]: "m a = None \<Longrightarrow> chg_map f a m = m"
wenzelm@46022
   212
  unfolding chg_map_def by auto
nipkow@19323
   213
wenzelm@46022
   214
lemma chg_map_upd[simp]: "m a = Some b \<Longrightarrow> chg_map f a m = m(a\<mapsto>f b)"
wenzelm@46022
   215
  unfolding chg_map_def by auto
nipkow@19323
   216
nipkow@19323
   217
lemma chg_map_other [simp]: "a \<noteq> b \<Longrightarrow> chg_map f a m b = m b"
wenzelm@46022
   218
  by (auto simp: chg_map_def)
nipkow@19323
   219
schirmer@12854
   220
schirmer@12854
   221
section "unique association lists"
schirmer@12854
   222
wenzelm@46022
   223
definition unique :: "('a \<times> 'b) list \<Rightarrow> bool"
wenzelm@38219
   224
  where "unique = distinct \<circ> map fst"
schirmer@12854
   225
wenzelm@46022
   226
lemma uniqueD: "unique l \<Longrightarrow> (x, y) \<in> set l \<Longrightarrow> (x', y') \<in> set l \<Longrightarrow> x = x' \<Longrightarrow> y = y'"
wenzelm@46022
   227
  unfolding unique_def o_def
wenzelm@46022
   228
  by (induct l) (auto dest: fst_in_set_lemma)
schirmer@12854
   229
schirmer@12854
   230
lemma unique_Nil [simp]: "unique []"
wenzelm@46022
   231
  by (simp add: unique_def)
schirmer@12854
   232
wenzelm@46022
   233
lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l \<and> (\<forall>y. (x,y) \<notin> set l))"
wenzelm@46022
   234
  by (auto simp: unique_def dest: fst_in_set_lemma)
schirmer@12854
   235
wenzelm@46022
   236
lemma unique_ConsD: "unique (x#xs) \<Longrightarrow> unique xs"
wenzelm@46022
   237
  by (simp add: unique_def)
schirmer@12854
   238
wenzelm@46022
   239
lemma unique_single [simp]: "\<And>p. unique [p]"
wenzelm@46022
   240
  by simp
schirmer@12854
   241
wenzelm@46022
   242
lemma unique_append [rule_format (no_asm)]: "unique l' \<Longrightarrow> unique l \<Longrightarrow>
wenzelm@46022
   243
    (\<forall>(x,y)\<in>set l. \<forall>(x',y')\<in>set l'. x' \<noteq> x) \<longrightarrow> unique (l @ l')"
wenzelm@46022
   244
  by (induct l) (auto dest: fst_in_set_lemma)
schirmer@12854
   245
wenzelm@46022
   246
lemma unique_map_inj: "unique l \<Longrightarrow> inj f \<Longrightarrow> unique (map (\<lambda>(k,x). (f k, g k x)) l)"
wenzelm@46022
   247
  by (induct l) (auto dest: fst_in_set_lemma simp add: inj_eq)
schirmer@12854
   248
wenzelm@46022
   249
lemma map_of_SomeI: "unique l \<Longrightarrow> (k, x) : set l \<Longrightarrow> map_of l k = Some x"
wenzelm@46022
   250
  by (induct l) auto
schirmer@12854
   251
schirmer@12854
   252
schirmer@12854
   253
section "list patterns"
schirmer@12854
   254
wenzelm@46022
   255
definition lsplit :: "[['a, 'a list] \<Rightarrow> 'b, 'a list] \<Rightarrow> 'b"
wenzelm@46022
   256
  where "lsplit = (\<lambda>f l. f (hd l) (tl l))"
wenzelm@38219
   257
wenzelm@38219
   258
text {* list patterns -- extends pre-defined type "pttrn" used in abstractions *}
schirmer@12854
   259
syntax
wenzelm@46022
   260
  "_lpttrn" :: "[pttrn, pttrn] \<Rightarrow> pttrn"    ("_#/_" [901,900] 900)
schirmer@12854
   261
translations
wenzelm@46022
   262
  "\<lambda>y # x # xs. b" \<rightleftharpoons> "CONST lsplit (\<lambda>y x # xs. b)"
wenzelm@46022
   263
  "\<lambda>x # xs. b" \<rightleftharpoons> "CONST lsplit (\<lambda>x xs. b)"
schirmer@12854
   264
schirmer@12854
   265
lemma lsplit [simp]: "lsplit c (x#xs) = c x xs"
wenzelm@46022
   266
  by (simp add: lsplit_def)
schirmer@12854
   267
schirmer@12854
   268
lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z"
wenzelm@46022
   269
  by (simp add: lsplit_def)
schirmer@12854
   270
schirmer@12854
   271
end