src/CCL/Hered.ML
author wenzelm
Sat, 17 Sep 2005 17:35:26 +0200
changeset 17456 bcf7544875b2
parent 5062 fbdb0b541314
permissions -rw-r--r--
converted to Isar theory format;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
     1
(*  Title:      CCL/Hered.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1459
d12da312eff4 expanded tabs
clasohm
parents: 757
diff changeset
     3
    Author:     Martin Coen, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
fun type_of_terms (Const("Trueprop",_) $ (Const("op =",(Type ("fun", [t,_])))$_$_)) = t;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
(*** Hereditary Termination ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    11
Goalw [HTTgen_def]  "mono(%X. HTTgen(X))";
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
    12
by (rtac monoI 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
by (fast_tac set_cs 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    14
qed "HTTgen_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    16
Goalw [HTTgen_def]
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
    17
  "t : HTTgen(A) <-> t=true | t=false | (EX a b. t=<a,b> & a : A & b : A) | \
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
    18
\                                       (EX f. t=lam x. f(x) & (ALL x. f(x) : A))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
by (fast_tac set_cs 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    20
qed "HTTgenXH";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    22
Goal
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
    23
  "t : HTT <-> t=true | t=false | (EX a b. t=<a,b> & a : HTT & b : HTT) | \
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
    24
\                                  (EX f. t=lam x. f(x) & (ALL x. f(x) : HTT))";
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
    25
by (rtac (rewrite_rule [HTTgen_def]
4423
a129b817b58a expandshort;
wenzelm
parents: 3837
diff changeset
    26
                 (HTTgen_mono RS (HTT_def RS def_gfp_Tarski) RS XHlemma1)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
by (fast_tac set_cs 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    28
qed "HTTXH";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
(*** Introduction Rules for HTT ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    32
Goal "~ bot : HTT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
by (fast_tac (term_cs addDs [XH_to_D HTTXH]) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    34
qed "HTT_bot";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    36
Goal "true : HTT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
by (fast_tac (term_cs addIs [XH_to_I HTTXH]) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    38
qed "HTT_true";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    40
Goal "false : HTT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
by (fast_tac (term_cs addIs [XH_to_I HTTXH]) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    42
qed "HTT_false";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    44
Goal "<a,b> : HTT <->  a : HTT  & b : HTT";
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
    45
by (rtac (HTTXH RS iff_trans) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
by (fast_tac term_cs 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    47
qed "HTT_pair";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
    49
Goal "lam x. f(x) : HTT <-> (ALL x. f(x) : HTT)";
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
    50
by (rtac (HTTXH RS iff_trans) 1);
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
    51
by (simp_tac term_ss 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
by (safe_tac term_cs);
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
    53
by (asm_simp_tac term_ss 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
by (fast_tac term_cs 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    55
qed "HTT_lam";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
local
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
  val raw_HTTrews = [HTT_bot,HTT_true,HTT_false,HTT_pair,HTT_lam];
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
    59
  fun mk_thm s = prove_goalw (the_context ()) data_defs s (fn _ =>
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
    60
                  [simp_tac (term_ss addsimps raw_HTTrews) 1]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
in
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
  val HTT_rews = raw_HTTrews @
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
               map mk_thm ["one : HTT",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
                           "inl(a) : HTT <-> a : HTT",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
                           "inr(b) : HTT <-> b : HTT",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
                           "zero : HTT",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
                           "succ(n) : HTT <-> n : HTT",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
                           "[] : HTT",
289
78541329ff35 changed "." to "$" to eliminate ambiguity
clasohm
parents: 8
diff changeset
    69
                           "x$xs : HTT <-> x : HTT & xs : HTT"];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
end;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
val HTT_Is = HTT_rews @ (HTT_rews RL [iffD2]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
(*** Coinduction for HTT ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    75
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
    76
val prems = goal (the_context ()) "[|  t : R;  R <= HTTgen(R) |] ==> t : HTT";
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
    77
by (rtac (HTT_def RS def_coinduct) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    78
by (REPEAT (ares_tac prems 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    79
qed "HTT_coinduct";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
fun HTT_coinduct_tac s i = res_inst_tac [("R",s)] HTT_coinduct i;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
    83
val prems = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
    "[|  t : R;   R <= HTTgen(lfp(%x. HTTgen(x) Un R Un HTT)) |] ==> t : HTT";
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
    85
by (rtac (HTTgen_mono RSN(3,HTT_def RS def_coinduct3)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
by (REPEAT (ares_tac prems 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
    87
qed "HTT_coinduct3";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
val HTT_coinduct3_raw = rewrite_rule [HTTgen_def] HTT_coinduct3;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
fun HTT_coinduct3_tac s i = res_inst_tac [("R",s)] HTT_coinduct3 i;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
    92
val HTTgenIs = map (mk_genIs (the_context ()) data_defs HTTgenXH HTTgen_mono)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
       ["true : HTTgen(R)",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
        "false : HTTgen(R)",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
        "[| a : R;  b : R |] ==> <a,b> : HTTgen(R)",
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
    96
        "[| !!x. b(x) : R |] ==> lam x. b(x) : HTTgen(R)",
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
        "one : HTTgen(R)",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
        "a : lfp(%x. HTTgen(x) Un R Un HTT) ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    99
\                         inl(a) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   100
        "b : lfp(%x. HTTgen(x) Un R Un HTT) ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
\                         inr(b) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   102
        "zero : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
        "n : lfp(%x. HTTgen(x) Un R Un HTT) ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
\                         succ(n) : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
        "[] : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
        "[| h : lfp(%x. HTTgen(x) Un R Un HTT); t : lfp(%x. HTTgen(x) Un R Un HTT) |] ==>\
289
78541329ff35 changed "." to "$" to eliminate ambiguity
clasohm
parents: 8
diff changeset
   107
\                         h$t : HTTgen(lfp(%x. HTTgen(x) Un R Un HTT))"];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   108
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
(*** Formation Rules for Types ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
   111
Goal "Unit <= HTT";
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
   112
by (simp_tac (CCL_ss addsimps ([subsetXH,UnitXH] @ HTT_rews)) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   113
qed "UnitF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   114
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
   115
Goal "Bool <= HTT";
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
   116
by (simp_tac (CCL_ss addsimps ([subsetXH,BoolXH] @ HTT_rews)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   118
qed "BoolF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   120
val prems = goal (the_context ()) "[| A <= HTT;  B <= HTT |] ==> A + B  <= HTT";
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
   121
by (simp_tac (CCL_ss addsimps ([subsetXH,PlusXH] @ HTT_rews)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   123
qed "PlusF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   125
val prems = goal (the_context ())
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
   126
     "[| A <= HTT;  !!x. x:A ==> B(x) <= HTT |] ==> SUM x:A. B(x) <= HTT";
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
   127
by (simp_tac (CCL_ss addsimps ([subsetXH,SgXH] @ HTT_rews)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   128
by (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD])) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   129
qed "SigmaF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
(*** Formation Rules for Recursive types - using coinduction these only need ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
(***                                          exhaution rule for type-former ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
(*Proof by induction - needs induction rule for type*)
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
   135
Goal "Nat <= HTT";
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
   136
by (simp_tac (term_ss addsimps [subsetXH]) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
by (safe_tac set_cs);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
   138
by (etac Nat_ind 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
by (ALLGOALS (fast_tac (set_cs addIs HTT_Is @ (prems RL [subsetD]))));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
val NatF = result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   141
5062
fbdb0b541314 isatool fixgoal;
wenzelm
parents: 4423
diff changeset
   142
Goal "Nat <= HTT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   143
by (safe_tac set_cs);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
   144
by (etac HTT_coinduct3 1);
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   145
by (fast_tac (set_cs addIs HTTgenIs
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   146
                 addSEs [HTTgen_mono RS ci3_RI] addEs [XH_to_E NatXH]) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   147
qed "NatF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   148
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   149
val [prem] = goal (the_context ()) "A <= HTT ==> List(A) <= HTT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   150
by (safe_tac set_cs);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
   151
by (etac HTT_coinduct3 1);
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   152
by (fast_tac (set_cs addSIs HTTgenIs
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   153
                 addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   154
                 addEs [XH_to_E ListXH]) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   155
qed "ListF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   156
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   157
val [prem] = goal (the_context ()) "A <= HTT ==> Lists(A) <= HTT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
by (safe_tac set_cs);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
   159
by (etac HTT_coinduct3 1);
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   160
by (fast_tac (set_cs addSIs HTTgenIs
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   161
                 addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   162
                 addEs [XH_to_E ListsXH]) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   163
qed "ListsF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   165
val [prem] = goal (the_context ()) "A <= HTT ==> ILists(A) <= HTT";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
by (safe_tac set_cs);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
   167
by (etac HTT_coinduct3 1);
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   168
by (fast_tac (set_cs addSIs HTTgenIs
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   169
                 addSEs [HTTgen_mono RS ci3_RI,prem RS subsetD RS (HTTgen_mono RS ci3_AI)]
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
                 addEs [XH_to_E IListsXH]) 1);
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   171
qed "IListsF";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   172
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
(*** A possible use for this predicate is proving equality from pre-order       ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   174
(*** but it seems as easy (and more general) to do this directly by coinduction ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
(*
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   176
val prems = goal (the_context ()) "[| t : HTT;  t [= u |] ==> u [= t";
3837
d7f033c74b38 fixed dots;
wenzelm
parents: 1459
diff changeset
   177
by (po_coinduct_tac "{p. EX a b. p=<a,b> & b : HTT & b [= a}" 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   178
by (fast_tac (ccl_cs addIs prems) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   179
by (safe_tac ccl_cs);
642
0db578095e6a CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents: 289
diff changeset
   180
by (dtac (poXH RS iffD1) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   181
by (safe_tac (set_cs addSEs [HTT_bot RS notE]));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   182
by (REPEAT_SOME (rtac (POgenXH RS iffD2) ORELSE' etac rev_mp));
8
c3d2c6dcf3f0 Installation of new simplfier. Previously appeared to set up the old
lcp
parents: 0
diff changeset
   183
by (ALLGOALS (simp_tac (term_ss addsimps HTT_rews)));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
by (ALLGOALS (fast_tac ccl_cs));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   185
qed "HTT_po_op";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   186
17456
bcf7544875b2 converted to Isar theory format;
wenzelm
parents: 5062
diff changeset
   187
val prems = goal (the_context ()) "[| t : HTT;  t [= u |] ==> t = u";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   188
by (REPEAT (ares_tac (prems @ [conjI RS (eq_iff RS iffD2),HTT_po_op]) 1));
757
2ca12511676d added qed and qed_goal[w]
clasohm
parents: 642
diff changeset
   189
qed "HTT_po_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
*)