| 12516 |      1 | (*  Title:      HOL/MicroJava/BV/Product.thy
 | 
| 10496 |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Tobias Nipkow
 | 
|  |      4 |     Copyright   2000 TUM
 | 
|  |      5 | 
 | 
|  |      6 | Products as semilattices
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
| 12911 |      9 | header {* \isaheader{Products as Semilattices} *}
 | 
| 10496 |     10 | 
 | 
| 27681 |     11 | theory Product
 | 
|  |     12 | imports Err
 | 
|  |     13 | begin
 | 
| 10496 |     14 | 
 | 
|  |     15 | constdefs
 | 
| 13006 |     16 |  le :: "'a ord \<Rightarrow> 'b ord \<Rightarrow> ('a * 'b) ord"
 | 
| 10496 |     17 | "le rA rB == %(a,b) (a',b'). a <=_rA a' & b <=_rB b'"
 | 
|  |     18 | 
 | 
| 13006 |     19 |  sup :: "'a ebinop \<Rightarrow> 'b ebinop \<Rightarrow> ('a * 'b)ebinop"
 | 
| 10496 |     20 | "sup f g == %(a1,b1)(a2,b2). Err.sup Pair (a1 +_f a2) (b1 +_g b2)"
 | 
|  |     21 | 
 | 
| 13006 |     22 |  esl :: "'a esl \<Rightarrow> 'b esl \<Rightarrow> ('a * 'b ) esl"
 | 
| 10496 |     23 | "esl == %(A,rA,fA) (B,rB,fB). (A <*> B, le rA rB, sup fA fB)"
 | 
|  |     24 | 
 | 
| 13006 |     25 | syntax "@lesubprod" :: "'a*'b \<Rightarrow> 'a ord \<Rightarrow> 'b ord \<Rightarrow> 'b \<Rightarrow> bool"
 | 
| 10496 |     26 |        ("(_ /<='(_,_') _)" [50, 0, 0, 51] 50)
 | 
|  |     27 | translations "p <=(rA,rB) q" == "p <=_(Product.le rA rB) q"
 | 
|  |     28 | 
 | 
|  |     29 | lemma unfold_lesub_prod:
 | 
|  |     30 |   "p <=(rA,rB) q == le rA rB p q"
 | 
|  |     31 |   by (simp add: lesub_def)
 | 
|  |     32 | 
 | 
|  |     33 | lemma le_prod_Pair_conv [iff]:
 | 
|  |     34 |   "((a1,b1) <=(rA,rB) (a2,b2)) = (a1 <=_rA a2 & b1 <=_rB b2)"
 | 
|  |     35 |   by (simp add: lesub_def le_def)
 | 
|  |     36 | 
 | 
|  |     37 | lemma less_prod_Pair_conv:
 | 
|  |     38 |   "((a1,b1) <_(Product.le rA rB) (a2,b2)) = 
 | 
|  |     39 |   (a1 <_rA a2 & b1 <=_rB b2 | a1 <=_rA a2 & b1 <_rB b2)"
 | 
|  |     40 | apply (unfold lesssub_def)
 | 
|  |     41 | apply simp
 | 
|  |     42 | apply blast
 | 
|  |     43 | done
 | 
|  |     44 | 
 | 
|  |     45 | lemma order_le_prod [iff]:
 | 
|  |     46 |   "order(Product.le rA rB) = (order rA & order rB)"
 | 
| 22068 |     47 | apply (unfold Semilat.order_def)
 | 
| 10496 |     48 | apply simp
 | 
| 15613 |     49 | apply blast
 | 
| 10496 |     50 | done 
 | 
|  |     51 | 
 | 
|  |     52 | 
 | 
|  |     53 | lemma acc_le_prodI [intro!]:
 | 
| 13006 |     54 |   "\<lbrakk> acc rA; acc rB \<rbrakk> \<Longrightarrow> acc(Product.le rA rB)"
 | 
| 10496 |     55 | apply (unfold acc_def)
 | 
| 22271 |     56 | apply (rule wfP_subset)
 | 
|  |     57 |  apply (erule wf_lex_prod [to_pred, THEN wfP_wf_eq [THEN iffD2]])
 | 
| 10496 |     58 |  apply assumption
 | 
|  |     59 | apply (auto simp add: lesssub_def less_prod_Pair_conv lex_prod_def)
 | 
|  |     60 | done
 | 
|  |     61 | 
 | 
|  |     62 | 
 | 
|  |     63 | lemma closed_lift2_sup:
 | 
| 13006 |     64 |   "\<lbrakk> closed (err A) (lift2 f); closed (err B) (lift2 g) \<rbrakk> \<Longrightarrow> 
 | 
| 10496 |     65 |   closed (err(A<*>B)) (lift2(sup f g))";
 | 
|  |     66 | apply (unfold closed_def plussub_def lift2_def err_def sup_def)
 | 
|  |     67 | apply (simp split: err.split)
 | 
|  |     68 | apply blast
 | 
|  |     69 | done 
 | 
|  |     70 | 
 | 
|  |     71 | lemma unfold_plussub_lift2:
 | 
|  |     72 |   "e1 +_(lift2 f) e2 == lift2 f e1 e2"
 | 
|  |     73 |   by (simp add: plussub_def)
 | 
|  |     74 | 
 | 
|  |     75 | 
 | 
|  |     76 | lemma plus_eq_Err_conv [simp]:
 | 
| 18372 |     77 |   assumes "x:A" and "y:A"
 | 
|  |     78 |     and "semilat(err A, Err.le r, lift2 f)"
 | 
|  |     79 |   shows "(x +_f y = Err) = (~(? z:A. x <=_r z & y <=_r z))"
 | 
| 10496 |     80 | proof -
 | 
|  |     81 |   have plus_le_conv2:
 | 
| 13006 |     82 |     "\<And>r f z. \<lbrakk> z : err A; semilat (err A, r, f); OK x : err A; OK y : err A;
 | 
|  |     83 |                  OK x +_f OK y <=_r z\<rbrakk> \<Longrightarrow> OK x <=_r z \<and> OK y <=_r z"
 | 
| 27681 |     84 |     by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1])
 | 
| 18372 |     85 |   from prems show ?thesis
 | 
| 10496 |     86 |   apply (rule_tac iffI)
 | 
|  |     87 |    apply clarify
 | 
|  |     88 |    apply (drule OK_le_err_OK [THEN iffD2])
 | 
|  |     89 |    apply (drule OK_le_err_OK [THEN iffD2])
 | 
| 27681 |     90 |    apply (drule Semilat.lub [OF Semilat.intro, of _ _ _ "OK x" _ "OK y"])
 | 
| 10496 |     91 |         apply assumption
 | 
|  |     92 |        apply assumption
 | 
|  |     93 |       apply simp
 | 
|  |     94 |      apply simp
 | 
|  |     95 |     apply simp
 | 
|  |     96 |    apply simp
 | 
|  |     97 |   apply (case_tac "x +_f y")
 | 
|  |     98 |    apply assumption
 | 
|  |     99 |   apply (rename_tac "z")
 | 
|  |    100 |   apply (subgoal_tac "OK z: err A")
 | 
|  |    101 |   apply (frule plus_le_conv2)
 | 
|  |    102 |        apply assumption
 | 
|  |    103 |       apply simp
 | 
|  |    104 |       apply blast
 | 
|  |    105 |      apply simp
 | 
| 27681 |    106 |     apply (blast dest: Semilat.orderI [OF Semilat.intro] order_refl)
 | 
| 10496 |    107 |    apply blast
 | 
|  |    108 |   apply (erule subst)
 | 
|  |    109 |   apply (unfold semilat_def err_def closed_def)
 | 
|  |    110 |   apply simp
 | 
|  |    111 |   done
 | 
|  |    112 | qed
 | 
|  |    113 | 
 | 
|  |    114 | lemma err_semilat_Product_esl:
 | 
| 13006 |    115 |   "\<And>L1 L2. \<lbrakk> err_semilat L1; err_semilat L2 \<rbrakk> \<Longrightarrow> err_semilat(Product.esl L1 L2)"
 | 
| 10496 |    116 | apply (unfold esl_def Err.sl_def)
 | 
|  |    117 | apply (simp (no_asm_simp) only: split_tupled_all)
 | 
|  |    118 | apply simp
 | 
|  |    119 | apply (simp (no_asm) only: semilat_Def)
 | 
| 27681 |    120 | apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup)
 | 
| 10496 |    121 | apply (simp (no_asm) only: unfold_lesub_err Err.le_def unfold_plussub_lift2 sup_def)
 | 
|  |    122 | apply (auto elim: semilat_le_err_OK1 semilat_le_err_OK2
 | 
|  |    123 |             simp add: lift2_def  split: err.split)
 | 
| 27681 |    124 | apply (blast dest: Semilat.orderI [OF Semilat.intro])
 | 
|  |    125 | apply (blast dest: Semilat.orderI [OF Semilat.intro])
 | 
| 10496 |    126 | 
 | 
|  |    127 | apply (rule OK_le_err_OK [THEN iffD1])
 | 
| 27681 |    128 | apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro])
 | 
| 10496 |    129 | apply simp
 | 
|  |    130 | apply simp
 | 
|  |    131 | apply simp
 | 
|  |    132 | apply simp
 | 
|  |    133 | apply simp
 | 
|  |    134 | apply simp
 | 
|  |    135 | 
 | 
|  |    136 | apply (rule OK_le_err_OK [THEN iffD1])
 | 
| 27681 |    137 | apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro])
 | 
| 10496 |    138 | apply simp
 | 
|  |    139 | apply simp
 | 
|  |    140 | apply simp
 | 
|  |    141 | apply simp
 | 
|  |    142 | apply simp
 | 
|  |    143 | apply simp
 | 
|  |    144 | done 
 | 
|  |    145 | 
 | 
|  |    146 | end
 |