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