src/ZF/ex/LList.ML
author wenzelm
Sun, 25 Oct 1998 12:33:27 +0100
changeset 5769 6a422b22ba02
parent 5529 4a54acae6a15
child 6141 a6922171b396
permissions -rw-r--r--
tuned checklist;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     1
(*  Title:      ZF/ex/LList.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
     4
    Copyright   1994  University of Cambridge
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
     6
Codatatype definition of Lazy Lists
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
     9
open LList;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    11
Delrules [subsetI, subsetCE];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    12
AddSIs [subset_refl, cons_subsetI, subset_consI, 
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    13
	Union_least, UN_least, Un_least, 
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    14
	Inter_greatest, Int_greatest, RepFun_subset,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    15
	Un_upper1, Un_upper2, Int_lower1, Int_lower2];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    16
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
(*An elimination rule, for type-checking*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    18
val LConsE = llist.mk_cases llist.con_defs "LCons(a,l) : llist(A)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
(*Proving freeness results*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    21
val LCons_iff      = llist.mk_free "LCons(a,l)=LCons(a',l') <-> a=a' & l=l'";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    22
val LNil_LCons_iff = llist.mk_free "~ LNil=LCons(a,l)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    24
Goal "llist(A) = {0} <+> (A <*> llist(A))";
529
f0d16216e394 ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
lcp
parents: 515
diff changeset
    25
let open llist;  val rew = rewrite_rule con_defs in  
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    26
by (fast_tac (claset() addSIs (subsetI ::map rew intrs) addEs [rew elim]) 1)
529
f0d16216e394 ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
lcp
parents: 515
diff changeset
    27
end;
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    28
qed "llist_unfold";
434
89d45187f04d Various updates and tidying
lcp
parents: 279
diff changeset
    29
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
(*** Lemmas to justify using "llist" in other recursive type definitions ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    32
Goalw llist.defs "A<=B ==> llist(A) <= llist(B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
by (rtac gfp_mono 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    34
by (REPEAT (rtac llist.bnd_mono 1));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
by (REPEAT (ares_tac (quniv_mono::basic_monos) 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    36
qed "llist_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
(** Closure of quniv(A) under llist -- why so complex?  Its a gfp... **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    40
AddSIs [QPair_Int_Vset_subset_UN RS subset_trans, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    41
                                 QPair_subset_univ,
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    42
                                 empty_subsetI, one_in_quniv RS qunivD];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    43
AddSDs [qunivD];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    44
AddSEs [Ord_in_Ord];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    46
Goal "Ord(i) ==> ALL l: llist(quniv(A)). l Int Vset(i) <= univ(eclose(A))";
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    47
by (etac trans_induct 1);
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
    48
by (rtac ballI 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    49
by (etac llist.elim 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    50
by (rewrite_goals_tac ([QInl_def,QInr_def]@llist.con_defs));
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    51
(*LNil case*)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    52
by (Asm_simp_tac 1);
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    53
(*LCons case*)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
    54
by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    55
qed "llist_quniv_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    57
Goal "llist(quniv(A)) <= quniv(A)";
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    58
by (rtac (qunivI RS subsetI) 1);
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    59
by (rtac Int_Vset_subset 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
by (REPEAT (ares_tac [llist_quniv_lemma RS bspec] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    61
qed "llist_quniv";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    62
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    63
bind_thm ("llist_subset_quniv",
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    64
    (llist_mono RS (llist_quniv RSN (2,subset_trans))));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    66
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    67
(*** Lazy List Equality: lleq ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    68
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    69
AddSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    70
AddSEs [Ord_in_Ord, Pair_inject];
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    71
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    72
(*Lemma for proving finality.  Unfold the lazy list; use induction hypothesis*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
    73
Goal "Ord(i) ==> ALL l l'. <l,l'> : lleq(A) --> l Int Vset(i) <= l'";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    74
by (etac trans_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    75
by (REPEAT (resolve_tac [allI, impI] 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    76
by (etac lleq.elim 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    77
by (rewrite_goals_tac (QInr_def::llist.con_defs));
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    78
by Safe_tac;
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    79
by (fast_tac (subset_cs addSEs [Ord_trans, make_elim bspec]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    80
qed "lleq_Int_Vset_subset_lemma";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    81
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    82
bind_thm ("lleq_Int_Vset_subset",
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
    83
        (lleq_Int_Vset_subset_lemma RS spec RS spec RS mp));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    84
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    85
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    86
(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    87
val [prem] = goal LList.thy "<l,l'> : lleq(A) ==> <l',l> : lleq(A)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    88
by (rtac (prem RS converseI RS lleq.coinduct) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    89
by (rtac (lleq.dom_subset RS converse_type) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    90
by Safe_tac;
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    91
by (etac lleq.elim 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    92
by (ALLGOALS Fast_tac);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    93
qed "lleq_symmetric";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    94
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
    95
Goal "<l,l'> : lleq(A) ==> l=l'";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    96
by (rtac equalityI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    97
by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
    98
     ORELSE etac lleq_symmetric 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    99
qed "lleq_implies_equal";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   100
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   101
val [eqprem,lprem] = goal LList.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   102
    "[| l=l';  l: llist(A) |] ==> <l,l'> : lleq(A)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   103
by (res_inst_tac [("X", "{<l,l>. l: llist(A)}")] lleq.coinduct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   104
by (rtac (lprem RS RepFunI RS (eqprem RS subst)) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   105
by Safe_tac;
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   106
by (etac llist.elim 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   107
by (ALLGOALS Fast_tac);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   108
qed "equal_llist_implies_leq";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   109
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   110
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   111
(*** Lazy List Functions ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   112
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   113
(*Examples of coinduction for type-checking and to prove llist equations*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   114
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   115
(*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   116
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   117
Goalw llist.con_defs "bnd_mono(univ(a), %l. LCons(a,l))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   118
by (rtac bnd_monoI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   119
by (REPEAT (ares_tac [subset_refl, QInr_mono, QPair_mono] 2));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   120
by (REPEAT (ares_tac [subset_refl, A_subset_univ, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 782
diff changeset
   121
                      QInr_subset_univ, QPair_subset_univ] 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   122
qed "lconst_fun_bnd_mono";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   123
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   124
(* lconst(a) = LCons(a,lconst(a)) *)
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   125
bind_thm ("lconst",
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   126
    ([lconst_def, lconst_fun_bnd_mono] MRS def_lfp_Tarski));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   127
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   128
val lconst_subset = lconst_def RS def_lfp_subset;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   129
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   130
bind_thm ("member_subset_Union_eclose", (arg_into_eclose RS Union_upper));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   131
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   132
Goal "a : A ==> lconst(a) : quniv(A)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   133
by (rtac (lconst_subset RS subset_trans RS qunivI) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   134
by (etac (arg_into_eclose RS eclose_subset RS univ_mono) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   135
qed "lconst_in_quniv";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   136
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   137
Goal "a:A ==> lconst(a): llist(A)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   138
by (rtac (singletonI RS llist.coinduct) 1);
576
469279790410 ZF/ex/LList/lconst_type: streamlined proof
lcp
parents: 529
diff changeset
   139
by (etac (lconst_in_quniv RS singleton_subsetI) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   140
by (fast_tac (claset() addSIs [lconst]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   141
qed "lconst_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   142
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   143
(*** flip --- equations merely assumed; certain consequences proved ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   144
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   145
Addsimps [flip_LNil, flip_LCons, not_type];
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   146
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   147
goal QUniv.thy "!!b. b:bool ==> b Int X <= univ(eclose(A))";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   148
by (fast_tac (claset() addIs [Int_lower1 RS subset_trans] addSEs [boolE]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   149
qed "bool_Int_subset_univ";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   150
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   151
AddSIs [not_type];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   152
AddIs  [bool_Int_subset_univ];
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   153
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   154
(*Reasoning borrowed from lleq.ML; a similar proof works for all
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   155
  "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5137
diff changeset
   156
Goal "Ord(i) ==> ALL l: llist(bool). flip(l) Int Vset(i) <= \
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   157
\                   univ(eclose(bool))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   158
by (etac trans_induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   159
by (rtac ballI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   160
by (etac llist.elim 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   161
by (ALLGOALS Asm_simp_tac);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   162
by (ALLGOALS 
5529
4a54acae6a15 tidying
paulson
parents: 5268
diff changeset
   163
    (asm_simp_tac (simpset() addsimps [QInl_def,QInr_def] @ llist.con_defs)));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   164
(*LCons case*)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   165
by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   166
qed "flip_llist_quniv_lemma";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   167
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   168
Goal "l: llist(bool) ==> flip(l) : quniv(bool)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   169
by (rtac (flip_llist_quniv_lemma RS bspec RS Int_Vset_subset RS qunivI) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   170
by (REPEAT (assume_tac 1));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   171
qed "flip_in_quniv";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   172
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   173
val [prem] = goal LList.thy "l : llist(bool) ==> flip(l): llist(bool)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   174
by (res_inst_tac [("X", "{flip(l) . l:llist(bool)}")]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   175
       llist.coinduct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   176
by (rtac (prem RS RepFunI) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   177
by (fast_tac (claset() addSIs [flip_in_quniv]) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   178
by (etac RepFunE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   179
by (etac llist.elim 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   180
by (ALLGOALS Asm_simp_tac);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   181
by (Fast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 576
diff changeset
   182
qed "flip_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   183
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   184
val [prem] = goal LList.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   185
    "l : llist(bool) ==> flip(flip(l)) = l";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   186
by (res_inst_tac [("X1", "{<flip(flip(l)),l> . l:llist(bool)}")]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   187
       (lleq.coinduct RS lleq_implies_equal) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   188
by (rtac (prem RS RepFunI) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   189
by (fast_tac (claset() addSIs [flip_type]) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   190
by (etac RepFunE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 486
diff changeset
   191
by (etac llist.elim 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   192
by (Asm_simp_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   193
by (asm_simp_tac (simpset() addsimps [flip_type, not_not]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2496
diff changeset
   194
by (fast_tac (claset() addSIs [not_type]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   195
qed "flip_flip";