src/HOL/BCV/Product.ML
author wenzelm
Thu, 12 Apr 2001 18:05:41 +0200
changeset 11253 caabb021ec0f
parent 10172 3daeda3d3cd0
permissions -rw-r--r--
proper order of order_less_asym';
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/BCV/Product.ML
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     2
    ID:         $Id$
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     4
    Copyright   2000 TUM
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     5
*)
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     6
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     7
Goalw [lesub_def] "p <=(rA,rB) q == le rA rB p q";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     8
by (Simp_tac 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
     9
qed "unfold_lesub_prod";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    10
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    11
Goalw [lesub_def,le_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    12
 "((a1,b1) <=(rA,rB) (a2,b2)) = (a1 <=_rA a2 & b1 <=_rB b2)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    13
by (Simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    14
qed "le_prod_Pair_conv";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    15
AddIffs [le_prod_Pair_conv];
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    16
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    17
Goalw [lesssub_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    18
 "((a1,b1) <_(Product.le rA rB) (a2,b2)) = \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    19
\ (a1 <_rA a2 & b1 <=_rB b2 | a1 <=_rA a2 & b1 <_rB b2)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    20
by (Simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    21
by (Blast_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    22
qed "less_prod_Pair_conv";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    23
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    24
Goalw [order_def] "order(Product.le rA rB) = (order rA & order rB)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    25
by (Simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    26
by (Blast_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    27
qed "order_le_prod";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    28
AddIffs [order_le_prod];
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    29
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    30
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    31
Goalw [acc_def] "[| acc rA; acc rB |] ==> acc(Product.le rA rB)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    32
by (rtac wf_subset 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    33
 by (etac wf_lex_prod 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    34
 by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    35
by (auto_tac (claset(), simpset()
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    36
     addsimps [lesssub_def,less_prod_Pair_conv,lex_prod_def]));
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    37
qed "acc_le_prodI";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    38
AddSIs [acc_le_prodI];
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    39
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    40
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    41
Goalw [closed_def,plussub_def,lift2_def,err_def,sup_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    42
 "[| closed (err A) (lift2 f); closed (err B) (lift2 g) |] ==> \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    43
\ closed (err(A<*>B)) (lift2(sup f g))";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    44
by (asm_full_simp_tac (simpset() addsplits [err.split]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    45
by (Blast_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    46
qed "closed_lift2_sup";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    47
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    48
Goalw [plussub_def] "e1 +_(lift2 f) e2 == lift2 f e1 e2";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    49
by (Simp_tac 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    50
qed "unfold_plussub_lift2";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    51
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    52
Goal
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    53
 "[| x:A; y:A; semilat(err A, Err.le r, lift2 f) |] ==> \
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    54
\ (x +_f y = Err) = (~(? z:A. x <=_r z & y <=_r z))";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    55
by (rtac iffI 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    56
 by (Clarify_tac 1);
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    57
 by (dtac (OK_le_err_OK RS iffD2) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    58
 by (dtac (OK_le_err_OK RS iffD2) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    59
 by (dtac semilat_lub 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    60
      by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    61
     by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    62
    by (Asm_simp_tac 1); 
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    63
   by (Asm_simp_tac 1); 
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    64
  by (Asm_simp_tac 1); 
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    65
 by (Asm_full_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    66
by (case_tac "x +_f y" 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    67
 by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    68
by (rename_tac "z" 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    69
by (subgoal_tac "OK z: err A" 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    70
 by (ftac (rotate_prems 2 (read_instantiate_sg(sign_of Err.thy)[("x","OK(x::'a)"),("y","OK(y::'a)")]plus_le_conv RS iffD1))1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    71
     by (assume_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    72
    by (Asm_simp_tac 1); 
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    73
    by (blast_tac (claset() addDs [semilatDorderI,order_refl]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    74
   by (Asm_simp_tac 1); 
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    75
  by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    76
 by (Blast_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    77
by (etac subst 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    78
by (rewrite_goals_tac [semilat_def,err_def,closed_def]);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    79
by (Asm_full_simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    80
qed "plus_eq_Err_conv";
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    81
Addsimps [plus_eq_Err_conv];
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    82
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    83
Goalw [esl_def,Err.sl_def]
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    84
 "!!L1 L2. [| err_semilat L1; err_semilat L2 |] ==> err_semilat(Product.esl L1 L2)";
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    85
by (split_all_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    86
by (Asm_full_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    87
by (simp_tac (HOL_basic_ss addsimps [semilat_Def]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    88
by (asm_simp_tac(HOL_basic_ss addsimps[semilatDclosedI,closed_lift2_sup])1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    89
by (simp_tac (HOL_basic_ss addsimps
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    90
     [unfold_lesub_err,Err.le_def,unfold_plussub_lift2,sup_def]) 1);
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    91
by (auto_tac (claset() addEs [semilat_le_err_OK1,semilat_le_err_OK2],
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    92
      simpset() addsimps [lift2_def] addsplits [err.split]));
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    93
by (blast_tac (claset() addDs [semilatDorderI]) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    94
by (blast_tac (claset() addDs [semilatDorderI]) 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
    95
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    96
by (rtac (OK_le_err_OK RS iffD1) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    97
by (etac subst 1 THEN stac (OK_lift2_OK RS sym) 1 THEN rtac semilat_lub 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    98
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
    99
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   100
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   101
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   102
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   103
by (Asm_simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   104
10172
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   105
by (rtac (OK_le_err_OK RS iffD1) 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   106
by (etac subst 1 THEN stac (OK_lift2_OK RS sym) 1 THEN rtac semilat_lub 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   107
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   108
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   109
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   110
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   111
by (Asm_simp_tac 1);
3daeda3d3cd0 expandshort
paulson
parents: 9791
diff changeset
   112
by (Asm_simp_tac 1);
9791
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   113
a39e5d43de55 Completely new version of BCV
nipkow
parents:
diff changeset
   114
qed "err_semilat_Product_esl";