src/HOL/Bali/Basis.thy
author haftmann
Tue, 07 Aug 2007 09:38:44 +0200
changeset 24162 8dfd5dd65d82
parent 24038 18182c4aec9e
child 24178 4ff1dc2aa18d
permissions -rw-r--r--
split off theory Option for benefit of code generator
     1 (*  Title:      HOL/Bali/Basis.thy
     2     ID:         $Id$
     3     Author:     David von Oheimb
     4 
     5 *)
     6 header {* Definitions extending HOL as logical basis of Bali *}
     7 
     8 theory Basis imports Main begin
     9 
    10 ML {*
    11 Unify.search_bound := 40;
    12 Unify.trace_bound  := 40;
    13 *}
    14 
    15 section "misc"
    16 
    17 declare same_fstI [intro!] (*### TO HOL/Wellfounded_Relations *)
    18 
    19 declare split_if_asm  [split] option.split [split] option.split_asm [split]
    20 declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
    21 declare if_weak_cong [cong del] option.weak_case_cong [cong del]
    22 declare length_Suc_conv [iff]
    23 
    24 lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}"
    25 apply auto
    26 done
    27 
    28 lemma subset_insertD: 
    29   "A <= insert x B ==> A <= B & x ~: A | (EX B'. A = insert x B' & B' <= B)"
    30 apply (case_tac "x:A")
    31 apply (rule disjI2)
    32 apply (rule_tac x = "A-{x}" in exI)
    33 apply fast+
    34 done
    35 
    36 syntax
    37   "3" :: nat   ("3") 
    38   "4" :: nat   ("4")
    39 translations
    40  "3" == "Suc 2"
    41  "4" == "Suc 3"
    42 
    43 (*unused*)
    44 lemma range_bool_domain: "range f = {f True, f False}"
    45 apply auto
    46 apply (case_tac "xa")
    47 apply auto
    48 done
    49 
    50 (* irrefl_tranclI in Transitive_Closure.thy is more general *)
    51 lemma irrefl_tranclI': "r^-1 Int r^+ = {} ==> !x. (x, x) ~: r^+"
    52 by(blast elim: tranclE dest: trancl_into_rtrancl)
    53 
    54 
    55 lemma trancl_rtrancl_trancl:
    56 "\<lbrakk>(x,y)\<in>r^+; (y,z)\<in>r^*\<rbrakk> \<Longrightarrow> (x,z)\<in>r^+"
    57 by (auto dest: tranclD rtrancl_trans rtrancl_into_trancl2)
    58 
    59 lemma rtrancl_into_trancl3:
    60 "\<lbrakk>(a,b)\<in>r^*; a\<noteq>b\<rbrakk> \<Longrightarrow> (a,b)\<in>r^+" 
    61 apply (drule rtranclD)
    62 apply auto
    63 done
    64 
    65 lemma rtrancl_into_rtrancl2: 
    66   "\<lbrakk> (a, b) \<in>  r; (b, c) \<in> r^* \<rbrakk> \<Longrightarrow> (a, c) \<in>  r^*"
    67 by (auto intro: r_into_rtrancl rtrancl_trans)
    68 
    69 lemma triangle_lemma:
    70  "\<lbrakk> \<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c; (a,x)\<in>r\<^sup>*; (a,y)\<in>r\<^sup>*\<rbrakk> 
    71  \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"
    72 proof -
    73   note converse_rtrancl_induct = converse_rtrancl_induct [consumes 1]
    74   note converse_rtranclE = converse_rtranclE [consumes 1] 
    75   assume unique: "\<And> a b c. \<lbrakk>(a,b)\<in>r; (a,c)\<in>r\<rbrakk> \<Longrightarrow> b=c"
    76   assume "(a,x)\<in>r\<^sup>*" 
    77   then show "(a,y)\<in>r\<^sup>* \<Longrightarrow> (x,y)\<in>r\<^sup>* \<or> (y,x)\<in>r\<^sup>*"
    78   proof (induct rule: converse_rtrancl_induct)
    79     assume "(x,y)\<in>r\<^sup>*"
    80     then show ?thesis 
    81       by blast
    82   next
    83     fix a v
    84     assume a_v_r: "(a, v) \<in> r" and
    85           v_x_rt: "(v, x) \<in> r\<^sup>*" and
    86           a_y_rt: "(a, y) \<in> r\<^sup>*"  and
    87              hyp: "(v, y) \<in> r\<^sup>* \<Longrightarrow> (x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
    88     from a_y_rt 
    89     show "(x, y) \<in> r\<^sup>* \<or> (y, x) \<in> r\<^sup>*"
    90     proof (cases rule: converse_rtranclE)
    91       assume "a=y"
    92       with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*"
    93 	by (auto intro: r_into_rtrancl rtrancl_trans)
    94       then show ?thesis 
    95 	by blast
    96     next
    97       fix w 
    98       assume a_w_r: "(a, w) \<in> r" and
    99             w_y_rt: "(w, y) \<in> r\<^sup>*"
   100       from a_v_r a_w_r unique 
   101       have "v=w" 
   102 	by auto
   103       with w_y_rt hyp 
   104       show ?thesis
   105 	by blast
   106     qed
   107   qed
   108 qed
   109 
   110 
   111 lemma rtrancl_cases [consumes 1, case_names Refl Trancl]:
   112  "\<lbrakk>(a,b)\<in>r\<^sup>*;  a = b \<Longrightarrow> P; (a,b)\<in>r\<^sup>+ \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
   113 apply (erule rtranclE)
   114 apply (auto dest: rtrancl_into_trancl1)
   115 done
   116 
   117 (* ### To Transitive_Closure *)
   118 theorems converse_rtrancl_induct 
   119  = converse_rtrancl_induct [consumes 1,case_names Id Step]
   120 
   121 theorems converse_trancl_induct 
   122          = converse_trancl_induct [consumes 1,case_names Single Step]
   123 
   124 (* context (theory "Set") *)
   125 lemma Ball_weaken:"\<lbrakk>Ball s P;\<And> x. P x\<longrightarrow>Q x\<rbrakk>\<Longrightarrow>Ball s Q"
   126 by auto
   127 
   128 (* context (theory "Finite") *)
   129 lemma finite_SetCompr2: "[| finite (Collect P); !y. P y --> finite (range (f y)) |] ==>  
   130   finite {f y x |x y. P y}"
   131 apply (subgoal_tac "{f y x |x y. P y} = UNION (Collect P) (%y. range (f y))")
   132 prefer 2 apply  fast
   133 apply (erule ssubst)
   134 apply (erule finite_UN_I)
   135 apply fast
   136 done
   137 
   138 
   139 (* ### TO theory "List" *)
   140 lemma list_all2_trans: "\<forall> a b c. P1 a b \<longrightarrow> P2 b c \<longrightarrow> P3 a c \<Longrightarrow>
   141  \<forall>xs2 xs3. list_all2 P1 xs1 xs2 \<longrightarrow> list_all2 P2 xs2 xs3 \<longrightarrow> list_all2 P3 xs1 xs3"
   142 apply (induct_tac "xs1")
   143 apply simp
   144 apply (rule allI)
   145 apply (induct_tac "xs2")
   146 apply simp
   147 apply (rule allI)
   148 apply (induct_tac "xs3")
   149 apply auto
   150 done
   151 
   152 
   153 section "pairs"
   154 
   155 lemma surjective_pairing5: "p = (fst p, fst (snd p), fst (snd (snd p)), fst (snd (snd (snd p))), 
   156   snd (snd (snd (snd p))))"
   157 apply auto
   158 done
   159 
   160 lemma fst_splitE [elim!]: 
   161 "[| fst s' = x';  !!x s. [| s' = (x,s);  x = x' |] ==> Q |] ==> Q"
   162 apply (cut_tac p = "s'" in surjective_pairing)
   163 apply auto
   164 done
   165 
   166 lemma fst_in_set_lemma [rule_format (no_asm)]: "(x, y) : set l --> x : fst ` set l"
   167 apply (induct_tac "l")
   168 apply  auto
   169 done
   170 
   171 
   172 section "quantifiers"
   173 
   174 lemma All_Ex_refl_eq2 [simp]: 
   175  "(!x. (? b. x = f b & Q b) \<longrightarrow> P x) = (!b. Q b --> P (f b))"
   176 apply auto
   177 done
   178 
   179 lemma ex_ex_miniscope1 [simp]:
   180   "(EX w v. P w v & Q v) = (EX v. (EX w. P w v) & Q v)"
   181 apply auto
   182 done
   183 
   184 lemma ex_miniscope2 [simp]:
   185   "(EX v. P v & Q & R v) = (Q & (EX v. P v & R v))" 
   186 apply auto
   187 done
   188 
   189 lemma ex_reorder31: "(\<exists>z x y. P x y z) = (\<exists>x y z. P x y z)"
   190 apply auto
   191 done
   192 
   193 lemma All_Ex_refl_eq1 [simp]: "(!x. (? b. x = f b) --> P x) = (!b. P (f b))"
   194 apply auto
   195 done
   196 
   197 
   198 section "sums"
   199 
   200 hide const In0 In1
   201 
   202 syntax
   203   fun_sum :: "('a => 'c) => ('b => 'c) => (('a+'b) => 'c)" (infixr "'(+')"80)
   204 translations
   205  "fun_sum" == "CONST sum_case"
   206 
   207 consts    the_Inl  :: "'a + 'b \<Rightarrow> 'a"
   208           the_Inr  :: "'a + 'b \<Rightarrow> 'b"
   209 primrec  "the_Inl (Inl a) = a"
   210 primrec  "the_Inr (Inr b) = b"
   211 
   212 datatype ('a, 'b, 'c) sum3 = In1 'a | In2 'b | In3 'c
   213 
   214 consts    the_In1  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'a"
   215           the_In2  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'b"
   216           the_In3  :: "('a, 'b, 'c) sum3 \<Rightarrow> 'c"
   217 primrec  "the_In1 (In1 a) = a"
   218 primrec  "the_In2 (In2 b) = b"
   219 primrec  "the_In3 (In3 c) = c"
   220 
   221 syntax
   222 	 In1l	:: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
   223 	 In1r	:: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
   224 translations
   225 	"In1l e" == "In1 (Inl e)"
   226 	"In1r c" == "In1 (Inr c)"
   227 
   228 syntax the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al"
   229        the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar"
   230 translations
   231    "the_In1l" == "the_Inl \<circ> the_In1"
   232    "the_In1r" == "the_Inr \<circ> the_In1"
   233 
   234 ML {*
   235 fun sum3_instantiate thm = map (fn s => simplify(simpset()delsimps[@{thm not_None_eq}])
   236  (read_instantiate [("t","In"^s^" ?x")] thm)) ["1l","2","3","1r"]
   237 *}
   238 (* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *)
   239 
   240 translations
   241   "option"<= (type) "Option.option"
   242   "list"  <= (type) "List.list"
   243   "sum3"  <= (type) "Basis.sum3"
   244 
   245 
   246 section "quantifiers for option type"
   247 
   248 syntax
   249   Oall :: "[pttrn, 'a option, bool] => bool"   ("(3! _:_:/ _)" [0,0,10] 10)
   250   Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3? _:_:/ _)" [0,0,10] 10)
   251 
   252 syntax (symbols)
   253   Oall :: "[pttrn, 'a option, bool] => bool"   ("(3\<forall>_\<in>_:/ _)"  [0,0,10] 10)
   254   Oex  :: "[pttrn, 'a option, bool] => bool"   ("(3\<exists>_\<in>_:/ _)"  [0,0,10] 10)
   255 
   256 translations
   257   "! x:A: P"    == "! x:o2s A. P"
   258   "? x:A: P"    == "? x:o2s A. P"
   259 
   260 section "Special map update"
   261 
   262 text{* Deemed too special for theory Map. *}
   263 
   264 constdefs
   265   chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)"
   266  "chg_map f a m == case m a of None => m | Some b => m(a|->f b)"
   267 
   268 lemma chg_map_new[simp]: "m a = None   ==> chg_map f a m = m"
   269 by (unfold chg_map_def, auto)
   270 
   271 lemma chg_map_upd[simp]: "m a = Some b ==> chg_map f a m = m(a|->f b)"
   272 by (unfold chg_map_def, auto)
   273 
   274 lemma chg_map_other [simp]: "a \<noteq> b \<Longrightarrow> chg_map f a m b = m b"
   275 by (auto simp: chg_map_def split add: option.split)
   276 
   277 
   278 section "unique association lists"
   279 
   280 constdefs
   281   unique   :: "('a \<times> 'b) list \<Rightarrow> bool"
   282  "unique \<equiv> distinct \<circ> map fst"
   283 
   284 lemma uniqueD [rule_format (no_asm)]: 
   285 "unique l--> (!x y. (x,y):set l --> (!x' y'. (x',y'):set l --> x=x'-->  y=y'))"
   286 apply (unfold unique_def o_def)
   287 apply (induct_tac "l")
   288 apply  (auto dest: fst_in_set_lemma)
   289 done
   290 
   291 lemma unique_Nil [simp]: "unique []"
   292 apply (unfold unique_def)
   293 apply (simp (no_asm))
   294 done
   295 
   296 lemma unique_Cons [simp]: "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))"
   297 apply (unfold unique_def)
   298 apply  (auto dest: fst_in_set_lemma)
   299 done
   300 
   301 lemmas unique_ConsI = conjI [THEN unique_Cons [THEN iffD2], standard]
   302 
   303 lemma unique_single [simp]: "!!p. unique [p]"
   304 apply auto
   305 done
   306 
   307 lemma unique_ConsD: "unique (x#xs) ==> unique xs"
   308 apply (simp add: unique_def)
   309 done
   310 
   311 lemma unique_append [rule_format (no_asm)]: "unique l' ==> unique l -->  
   312   (!(x,y):set l. !(x',y'):set l'. x' ~= x) --> unique (l @ l')"
   313 apply (induct_tac "l")
   314 apply  (auto dest: fst_in_set_lemma)
   315 done
   316 
   317 lemma unique_map_inj [rule_format (no_asm)]: "unique l --> inj f --> unique (map (%(k,x). (f k, g k x)) l)"
   318 apply (induct_tac "l")
   319 apply  (auto dest: fst_in_set_lemma simp add: inj_eq)
   320 done
   321 
   322 lemma map_of_SomeI [rule_format (no_asm)]: "unique l --> (k, x) : set l --> map_of l k = Some x"
   323 apply (induct_tac "l")
   324 apply auto
   325 done
   326 
   327 
   328 section "list patterns"
   329 
   330 consts
   331   lsplit         :: "[['a, 'a list] => 'b, 'a list] => 'b"
   332 defs
   333   lsplit_def:    "lsplit == %f l. f (hd l) (tl l)"
   334 (*  list patterns -- extends pre-defined type "pttrn" used in abstractions *)
   335 syntax
   336   "_lpttrn"    :: "[pttrn,pttrn] => pttrn"     ("_#/_" [901,900] 900)
   337 translations
   338   "%y#x#xs. b"  == "lsplit (%y x#xs. b)"
   339   "%x#xs  . b"  == "lsplit (%x xs  . b)"
   340 
   341 lemma lsplit [simp]: "lsplit c (x#xs) = c x xs"
   342 apply (unfold lsplit_def)
   343 apply (simp (no_asm))
   344 done
   345 
   346 lemma lsplit2 [simp]: "lsplit P (x#xs) y z = P x xs y z"
   347 apply (unfold lsplit_def)
   348 apply simp
   349 done 
   350 
   351 
   352 end