src/HOL/BCV/Orders0.ML
author nipkow
Mon, 13 Mar 2000 12:51:10 +0100
changeset 8423 3c19160b6432
parent 7961 422ac6888c7f
child 8442 96023903c2df
permissions -rw-r--r--
exhaust_tac -> cases_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/BCV/Orders0.ML
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     4
    Copyright   1999 TUM
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     5
*)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     6
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     7
(** option **)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     8
section "option";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
     9
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    10
Goalw [le_option] "(o1::('a::order)option) <= o1";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    11
by (simp_tac (simpset() addsplits [option.split]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    12
qed "le_option_refl";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    13
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    14
Goalw [le_option] "(o1::('a::order)option) <= o2 --> o2<=o3 --> o1<=o3";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    15
by (simp_tac (simpset() addsplits [option.split]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    16
by (blast_tac (claset() addIs [order_trans]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    17
qed_spec_mp "le_option_trans";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    18
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    19
Goalw [le_option] "(o1::('a::order)option) <= o2 --> o2<=o1 --> o1=o2";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    20
by (simp_tac (simpset() addsplits [option.split]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    21
by (blast_tac (claset() addIs [order_antisym]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    22
qed_spec_mp "le_option_antisym";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    23
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    24
Goalw [less_option] "(o1::('a::ord)option) < o2 = (o1<=o2 & o1 ~= o2)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    25
by (rtac refl 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    26
qed "less_le_option";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    27
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    28
Goalw [le_option] "Some x <= opt = (? y. opt = Some y & x <= y)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    29
by (simp_tac (simpset() addsplits [option.split]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    30
qed_spec_mp "Some_le_conv";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    31
AddIffs [Some_le_conv];
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    32
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    33
Goalw [le_option] "None <= opt";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    34
by (simp_tac (simpset() addsplits [option.split]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    35
qed "None_le";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    36
AddIffs [None_le];
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    37
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    38
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    39
(** list **)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    40
section "list";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    41
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    42
Goalw [le_list] "[| xs <= ys; p < size xs |] ==> xs!p <= ys!p";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    43
by (Blast_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    44
qed "le_listD";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    45
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    46
Goalw [le_list] "([] <= ys) = (ys = [])";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    47
by (Simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    48
qed "Nil_le_conv";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    49
AddIffs [Nil_le_conv];
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    50
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    51
Goalw [le_list] "(xs::('a::order)list) <= xs";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    52
by (induct_tac "xs" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    53
by (Auto_tac);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    54
qed "le_list_refl";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    55
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    56
Goalw [le_list]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    57
 "!ys zs.(xs::('a::order)list) <= ys --> ys <= zs --> xs <= zs";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    58
by (induct_tac "xs" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    59
 by (Simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    60
by (rtac allI 1);
8423
3c19160b6432 exhaust_tac -> cases_tac
nipkow
parents: 7961
diff changeset
    61
by (cases_tac "ys" 1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    62
 by (hyp_subst_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    63
 by (Simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    64
by (rtac allI 1);
8423
3c19160b6432 exhaust_tac -> cases_tac
nipkow
parents: 7961
diff changeset
    65
by (cases_tac "zs" 1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    66
 by (hyp_subst_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    67
 by (Simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    68
by (hyp_subst_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    69
by (simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    70
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    71
by (rtac conjI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    72
 by (REPEAT(EVERY1[etac allE, etac conjE, etac impE, rtac refl]));
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    73
 by (blast_tac (claset() addIs [order_trans]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    74
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    75
by (EVERY[etac allE 1, etac allE 1, etac impE 1, etac impE 2]);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    76
  by (rtac conjI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    77
  by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    78
  by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    79
 by (rtac conjI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    80
 by (assume_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    81
 by (Blast_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    82
by (Asm_full_simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    83
qed_spec_mp "le_list_trans";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    84
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    85
Goalw [le_list]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    86
 "!ys. (xs::('a::order)list) <= ys --> ys <= xs --> xs = ys";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    87
by (induct_tac "xs" 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    88
 by (Simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    89
by (rtac allI 1);
8423
3c19160b6432 exhaust_tac -> cases_tac
nipkow
parents: 7961
diff changeset
    90
by (cases_tac "ys" 1);
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    91
 by (hyp_subst_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    92
 by (Simp_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    93
by (hyp_subst_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    94
by (simp_tac (simpset() addsimps [nth_Cons] addsplits [nat.split]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    95
by (Clarify_tac 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    96
by (rtac conjI 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    97
 by (blast_tac (claset() addIs [order_antisym]) 1);
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
    98
by (Asm_full_simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
    99
qed_spec_mp "le_list_antisym";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   100
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   101
Goalw [less_list] "(xs::('a::ord)list) < ys = (xs<=ys & xs ~= ys)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   102
by (rtac refl 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   103
qed "less_le_list";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   104
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   105
(** product **)
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   106
section "product";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   107
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   108
Goalw [le_prod] "(p1::('a::order * 'b::order)) <= p1";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   109
by (Simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   110
qed "le_prod_refl";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   111
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   112
Goalw [le_prod]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   113
 "[| (p1::('a::order * 'b::order)) <= p2; p2<=p3 |] ==> p1<=p3";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   114
by (blast_tac (claset() addIs [order_trans]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   115
qed "le_prod_trans";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   116
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   117
Goalw [le_prod]
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   118
 "[| (p1::('a::order * 'b::order)) <= p2; p2 <= p1 |] ==> p1 = p2";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   119
by (blast_tac (claset() addIs [order_antisym,trans,surjective_pairing,sym]) 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   120
qed_spec_mp "le_prod_antisym";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   121
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   122
Goalw [less_prod] "(p1::('a::order * 'b::order)) < p2 = (p1<=p2 & p1 ~= p2)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   123
by (rtac refl 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   124
qed "less_le_prod";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   125
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   126
Goalw [le_prod] "((a,b) <= (c,d)) = (a <= c & b <= d)";
7961
422ac6888c7f expandshort
paulson
parents: 7626
diff changeset
   127
by (Simp_tac 1);
7626
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   128
qed "le_prod_iff";
5997f35954d7 A new theory: a model of bytecode verification.
nipkow
parents:
diff changeset
   129
AddIffs [le_prod_iff];