src/HOLCF/Lift3.ML
author nipkow
Wed, 26 Mar 1997 17:58:48 +0100
changeset 2841 c2508f4ab739
parent 2648 9944bea3b459
child 3035 5230c37ad29f
permissions -rw-r--r--
Added "discrete" CPOs and modified IMP to use those rather than "lift"
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2357
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     1
(*  Title:      HOLCF/Lift3.ML
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     2
    ID:         $Id$
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     3
    Author:     Olaf Mueller, Robert Sandner
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     4
    Copyright   1996 Technische Universitaet Muenchen
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     5
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     6
Theorems for Lift3.thy
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     7
*)
dd2e5e655fd2 Headers added
sandnerr
parents: 2356
diff changeset
     8
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
     9
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    10
(* for compatibility with old HOLCF-Version *)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    11
qed_goal "inst_lift_pcpo" thy "UU = Undef"
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    12
 (fn prems => 
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    13
        [
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    14
        (simp_tac (HOL_ss addsimps [UU_def,UU_lift_def]) 1)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    15
        ]);
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
    16
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    17
(* ----------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    18
(*                        From Undef to UU		       *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    19
(* ----------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    20
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    21
Addsimps [inst_lift_pcpo];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    22
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    23
local
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    24
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    25
val case1' = prove_goal Lift3.thy "lift_case f1 f2 UU = f1"
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    26
	     (fn _ => [simp_tac (!simpset addsimps lift.simps) 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    27
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    28
val case2' = prove_goal Lift3.thy "lift_case f1 f2 (Def a) = f2 a"
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    29
	     (fn _ => [Simp_tac 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    30
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    31
val distinct1' = prove_goal Lift3.thy "UU ~= Def a" 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    32
		 (fn _ => [Simp_tac 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    33
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    34
val distinct2' = prove_goal Lift3.thy "Def a ~= UU"
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    35
		 (fn _ => [Simp_tac 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    36
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    37
val inject' = prove_goal Lift3.thy "Def a = Def aa = (a = aa)"
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    38
	       (fn _ => [Simp_tac 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    39
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    40
val rec1' = prove_goal Lift3.thy "lift_rec f1 f2 UU = f1"
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    41
	    (fn _ => [Simp_tac 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    42
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    43
val rec2' = prove_goal Lift3.thy "lift_rec f1 f2 (Def a) = f2 a"
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    44
	    (fn _ => [Simp_tac 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    45
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    46
val induct' = prove_goal Lift3.thy "[| P UU; !a. P (Def a) |] ==> P lift"
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    47
	    (fn prems => [cut_facts_tac prems 1, Asm_full_simp_tac 1,
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    48
		      etac Lift1.lift.induct 1,fast_tac HOL_cs 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    49
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    50
in 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    51
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    52
val Def_not_UU = distinct1' RS not_sym;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    53
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    54
structure lift =
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    55
struct
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    56
val cases = [case1',case2'];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    57
val distinct = [distinct1',distinct2'];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    58
val inject = [inject'];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    59
val induct = allI RSN(2,induct');
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    60
val recs = [rec1',rec2'];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    61
val simps = cases@distinct@inject@recs;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    62
fun induct_tac (s:string) (i:int) = 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    63
    (res_inst_tac [("lift",s)] induct i);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    64
end;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    65
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    66
end; (* local *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    67
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    68
Delsimps [inst_lift_pcpo];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    69
Delsimps lift.simps;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    70
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    71
Addsimps [inst_lift_pcpo RS sym];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    72
Addsimps lift.simps;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    73
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    74
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    75
(* -------------------------------------------------------------------------*)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    76
(* rewrite_rule for less_lift 						    *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    77
(* -------------------------------------------------------------------------*)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    78
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    79
goal Lift3.thy "(x::'a lift) << y = (x=y | x=UU)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    80
br (inst_lift_po RS ssubst) 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    81
by (Simp_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    82
val less_lift = result();
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    83
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    84
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    85
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    86
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    87
(* ---------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    88
(*                  Relating UU and Undef                     *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    89
(* ---------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    90
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    91
goal Lift3.thy "x=UU | (? y.x=Def y)"; 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    92
by (lift.induct_tac "x" 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    93
by (Asm_simp_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    94
by (rtac disjI2 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    95
by (rtac exI 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    96
by (Asm_simp_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    97
qed"Lift_exhaust";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    98
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
    99
val prems = goal Lift3.thy 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   100
  "[| x = UU ==> P; ? a. x = Def a ==> P |] ==> P";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   101
by (cut_facts_tac [Lift_exhaust] 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   102
by (fast_tac (HOL_cs addSEs prems) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   103
qed"Lift_cases";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   104
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2648
diff changeset
   105
goal thy
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2648
diff changeset
   106
  "P(lift_case a b x) = ((x=UU --> P a) & (!y. x = Def y --> P(b y)))";
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2648
diff changeset
   107
by(lift.induct_tac "x" 1);
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2648
diff changeset
   108
by(ALLGOALS Asm_simp_tac);
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2648
diff changeset
   109
qed "expand_lift_case";
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2648
diff changeset
   110
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   111
goal Lift3.thy "(x~=UU)=(? y.x=Def y)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   112
br iffI 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   113
 br Lift_cases 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   114
  by (fast_tac HOL_cs 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   115
 by (fast_tac HOL_cs 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   116
by (fast_tac (HOL_cs addSIs lift.distinct) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   117
qed"not_Undef_is_Def";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   118
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   119
val Undef_eq_UU = inst_lift_pcpo RS sym;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   120
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   121
val DefE = prove_goal Lift3.thy "Def x = UU ==> R" 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   122
    (fn prems => [
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   123
        cut_facts_tac prems 1,
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   124
        asm_full_simp_tac (HOL_ss addsimps [Def_not_UU]) 1]);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   125
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   126
val prems = goal Lift3.thy "[| x = Def s; x = UU |] ==> R";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   127
by (cut_facts_tac prems 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   128
by (fast_tac (HOL_cs addSDs [DefE]) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   129
val DefE2 = result();
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   130
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   131
(* ---------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   132
(*                          Lift is flat                     *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   133
(* ---------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   134
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   135
goalw Lift3.thy [flat_def] "flat (x::'a lift)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   136
by (simp_tac (!simpset addsimps [less_lift]) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   137
val flat_lift = result();
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   138
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2566
diff changeset
   139
bind_thm("ax_flat_lift",flat_lift RS flatE);
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   140
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   141
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   142
(* ---------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   143
(*       More Continuity Proofs and Extended Tactic           *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   144
(* ---------------------------------------------------------- *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   145
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   146
goal Lift3.thy "cont (%x. case x of Undef => UU | Def a => f a)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   147
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   148
br flatdom_strict2cont 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   149
 br flat_lift 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   150
by (Simp_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   151
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   152
val cont_flift1_arg = result();
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   153
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   154
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   155
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   156
goal Lift3.thy "cont (%x. case x of Undef => UU | Def a => Def (f a))";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   157
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   158
br flatdom_strict2cont 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   159
 br flat_lift 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   160
by (Simp_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   161
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   162
val cont_flift2_arg = result();
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   163
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   164
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   165
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   166
goal Lift3.thy "!!f. [|! a.cont (%y. (f y) a)|] ==> \
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   167
\   cont (%y. case x of Undef => UU | Def a => (f y) a)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   168
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   169
by (res_inst_tac [("x","x")] Lift_cases 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   170
 by (Asm_simp_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   171
by (fast_tac (HOL_cs addss !simpset) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   172
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   173
qed"cont_flift1_not_arg";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   174
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   175
val cont_flift1_not_arg2 = (allI RS cont_flift1_not_arg);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   176
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   177
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   178
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   179
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   180
(* zusammenfassen zu cont(%y. ((f y)`(g y)) s)     *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   181
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   182
goal Lift3.thy "!!f.cont g ==> cont(%x. (f`(g x)) s)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   183
by (rtac monocontlub2cont 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   184
(* monotone *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   185
 by (rtac monofunI 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   186
 by (strip_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   187
 by (rtac (monofun_cfun_arg RS monofun_fun_fun) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   188
 by (etac (cont2mono RS monofunE RS spec RS spec RS mp) 1); 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   189
 by (atac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   190
(* contlub *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   191
by (rtac contlubI 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   192
by (strip_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   193
by ((rtac (cont2contlub RS contlubE RS spec RS mp RS ssubst) 1) THEN (atac 1));
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   194
 ba 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   195
by (stac (contlub_cfun_arg RS fun_cong) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   196
 be (cont2mono RS ch2ch_monofun) 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   197
 ba 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   198
by (stac thelub_fun 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   199
 by (fast_tac ((HOL_cs addSIs [ch2ch_fappR]) 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   200
		       addSEs [cont2mono RS ch2ch_monofun]) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   201
br refl 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   202
qed"cont_fapp_app1";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   203
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   204
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   205
goal Lift3.thy "cont(%y. (y`x) s)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   206
by (rtac monocontlub2cont 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   207
 (* monotone *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   208
 by (rtac monofunI 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   209
 by (strip_tac 1); 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   210
 be (monofun_cfun_fun RS monofun_fun_fun) 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   211
(* continuous *)
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   212
by (rtac contlubI 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   213
by (strip_tac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   214
by (stac (contlub_cfun_fun RS fun_cong) 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   215
 by (atac 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   216
by (stac thelub_fun 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   217
 be ch2ch_fappL 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   218
br refl 1;
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   219
qed"cont_fapp_app2";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   220
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   221
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   222
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   223
val prems = goal Lift3.thy "[| cont f1; cont f2 |] ==> \
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   224
\   cont (%x. if b then f1 x else f2 x)";
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   225
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   226
by (cut_facts_tac prems 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   227
by (case_tac "b" 1);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   228
by (TRYALL (fast_tac (HOL_cs addss HOL_ss)));
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   229
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   230
val cont_if = result();
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   231
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   232
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   233
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   234
val cont2cont_CF1L_rev2 = (allI RS cont2cont_CF1L_rev);
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   235
2566
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2357
diff changeset
   236
val cont_lemmas2 =  cont_lemmas1@
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   237
		   [cont_flift1_arg,cont_flift2_arg,
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   238
                    cont_flift1_not_arg2,cont2cont_CF1L_rev2, 
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   239
                    cont_fapp_app1,cont_fapp_app2,cont_if];
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   240
2566
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2357
diff changeset
   241
Addsimps 	   [cont_flift1_arg,cont_flift2_arg,
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2357
diff changeset
   242
                    cont_flift1_not_arg2,cont2cont_CF1L_rev2, 
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2357
diff changeset
   243
                    cont_fapp_app1,cont_fapp_app2,cont_if];
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   244
2566
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2357
diff changeset
   245
fun cont_tac  i = resolve_tac cont_lemmas2 i;
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2357
diff changeset
   246
fun cont_tacR i = simp_tac (!simpset addsimps [flift1_def,flift2_def]) i THEN
cbf02fc74332 changed handling of cont_lemmas and adm_lemmas
oheimb
parents: 2357
diff changeset
   247
		  REPEAT (cont_tac i);
2356
125260ef480c Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
sandnerr
parents:
diff changeset
   248
2646
099a9155f608 reflecting recent changes of the simplifier
oheimb
parents: 2640
diff changeset
   249
simpset := !simpset addSolver (K (DEPTH_SOLVE_1 o cont_tac));