src/HOLCF/ex/Witness.ML
author wenzelm
Fri, 07 Mar 1997 11:48:46 +0100
changeset 2754 59bd96046ad6
parent 2642 3c3a84cc85a9
child 2932 9c4d5fd41c9b
permissions -rw-r--r--
moved settings comment to build;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     1
open Witness;
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     2
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     3
(* -------------------------------------------------------------------- *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     4
(* classes cplus, cminus, ctimes, cdiv introduce 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     5
   characteristic constants  o+ o- o* o/
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     6
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
     7
	"circ":: "one -> one -> one"
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     8
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     9
   is the witness for o+ o- o* o/	
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    10
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    11
   No characteristic axioms are to be checked
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    12
*)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    13
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    14
(* -------------------------------------------------------------------- *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    15
(* classes per and qpo introduce characteristic constants
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    16
	".="	:: "'a::per -> 'a -> tr"		(cinfixl 55)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    17
	".<="	:: "'a::qpo -> 'a -> tr"		(cinfixl 55)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    18
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    19
   The witness for these is 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    20
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    21
	"bullet":: "one -> one -> tr"			(cinfixl 55)
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    22
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    23
   the classes equiv, eq, impose additional axioms
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    24
*)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    25
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    26
(* -------------------------------------------------------------------- *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    27
(* 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    28
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    29
characteristic axioms of class per:
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    30
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    31
strict_per	"(UU .= x) = UU & (x .= UU) = UU"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    32
total_per	"[|x ~= UU; y ~= UU|] ==> (x .= y) ~= UU"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    33
flat_per	"flat (UU::'a::per)"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    34
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    35
sym_per		"(x .= y) = (y .= x)"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    36
trans_per	"[|(x .= y)=TT; (y .= z)=TT |] ==> (x .= z)=TT"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    37
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    38
   --------------------------------------------------------------------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    39
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    40
characteristic axioms of class equiv:
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    41
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    42
refl_per	"[|(x::'a::equiv) ~= UU|] ==> (x .= x)=TT"
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    43
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    44
   --------------------------------------------------------------------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    45
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    46
characteristic axioms of class eq:
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    47
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    48
weq	"[|(x::'a::eq)~=UU; y~=UU|] ==> (x=y --> (x.=y)=TT) & (x~=y --> Çx.=yÈ)"
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    49
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    50
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    51
   --------------------------------------------------------------------
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    52
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    53
*)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    54
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    55
(* strict_per, strict_qpo *)
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    56
goalw thy [bullet_def] "(UU bullet x) = UU & (x bullet UU) = UU";
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    57
by (simp_tac (!simpset addsimps [flift1_def,flift2_def,o_def]) 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    58
by (lift.induct_tac "x" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    59
auto();
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    60
result();
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    61
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    62
(* total_per, total_qpo *)
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    63
val prems = goal thy "[|x~=UU; y~=UU|] ==> (x bullet y) ~= UU";
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    64
by (subgoal_tac "x~=UU&y~=UU-->(x bullet y) ~= UU" 1);
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    65
by (cut_facts_tac prems 1);
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    66
by (fast_tac HOL_cs 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    67
by (simp_tac (!simpset addsimps [bullet_def,flift1_def,flift2_def,o_def]) 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    68
by (lift.induct_tac "x" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    69
by (fast_tac HOL_cs 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    70
by (lift.induct_tac "y" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    71
auto();
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    72
result();
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    73
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    74
(* flat_per *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    75
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    76
goal thy "flat (x::one)";
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    77
by (rtac flat_flat 1);
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    78
result();
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    79
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    80
(* sym_per *)  
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    81
goalw thy [bullet_def] "(x bullet y) = (y bullet x)";
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    82
by (simp_tac (!simpset addsimps [bullet_def,flift1_def,flift2_def,o_def]) 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    83
by (lift.induct_tac "x" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    84
by (lift.induct_tac "y" 2);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    85
by (lift.induct_tac "y" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    86
auto();
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    87
result();
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    88
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    89
(* trans_per, trans_qpo *)
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    90
val prems = goal thy "[|(x bullet y)=TT; (y bullet z)=TT|] ==>(x bullet z)=TT";
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    91
by (subgoal_tac "(x bullet y)=TT&(y bullet z)=TT-->(x bullet z)=TT" 1);
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    92
by (cut_facts_tac prems 1);
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    93
by (fast_tac HOL_cs 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    94
by (simp_tac (!simpset addsimps [bullet_def,flift1_def,flift2_def,o_def]) 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    95
by (lift.induct_tac "x" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    96
by (lift.induct_tac "y" 2);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    97
by (lift.induct_tac "y" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    98
by (lift.induct_tac "z" 4);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
    99
by (lift.induct_tac "z" 3);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   100
by (lift.induct_tac "z" 2);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   101
by (lift.induct_tac "z" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   102
auto();
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   103
result();
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   104
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   105
(* refl_per *)
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   106
val prems = goal thy "x ~= UU ==> (x bullet x)=TT";
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   107
by (subgoal_tac "x ~= UU --> (x bullet x)=TT" 1);
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   108
by (cut_facts_tac prems 1);
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   109
by (fast_tac HOL_cs 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   110
by (simp_tac (!simpset addsimps [bullet_def,flift1_def,flift2_def,o_def]) 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   111
by (lift.induct_tac "x" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   112
auto();
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   113
qed "refl_per_one";
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   114
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   115
(* weq *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   116
val prems = goal thy 
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   117
	"[|x~=UU; y~=UU|]==>(x=y-->(x bullet y)=TT)&(x~=y-->(x bullet y)=FF)";
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   118
by(subgoal_tac"x~=UU&y~=UU-->(x=y-->(x bullet y)=TT)&(x~=y-->(x bullet y)=FF)"1);
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   119
by (cut_facts_tac prems 1);
2642
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   120
by (fast_tac HOL_cs 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   121
by (lift.induct_tac "x" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   122
by (lift.induct_tac "y" 2);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   123
by (lift.induct_tac "y" 1);
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   124
auto();
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   125
br refl_per_one 1;
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   126
auto();
3c3a84cc85a9 Examples are adopted to the changes from HOLCF.
slotosch
parents: 2570
diff changeset
   127
by (simp_tac (!simpset addsimps [bullet_def,flift1_def,flift2_def,o_def]) 1);
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   128
result();
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
   129