src/ZF/ex/LList.thy
author wenzelm
Thu, 26 Jul 2012 11:46:30 +0200
changeset 48507 1182677e4728
parent 46935 38ecb2dc3636
child 57492 74bf65a1910a
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/ex/LList.thy
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     3
    Copyright   1994  University of Cambridge
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
     5
Codatatype definition of Lazy Lists.
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     6
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
Equality for llist(A) as a greatest fixed point
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
Functions for Lazy Lists
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
STILL NEEDS:
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
co_recursion for defining lconst, flip, etc.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
a typing rule for it, based on some notion of "productivity..."
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    15
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13339
diff changeset
    16
theory LList imports Main begin
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    17
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    18
consts
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    19
  llist  :: "i=>i";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    20
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    21
codatatype
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    22
  "llist(A)" = LNil | LCons ("a \<in> A", "l \<in> llist(A)")
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    23
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    24
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    25
(*Coinductive definition of equality*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    26
consts
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    27
  lleq :: "i=>i"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    28
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
(*Previously used <*> in the domain and variant pairs as elements.  But
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
  standard pairs work just as well.  To use variant pairs, must change prefix
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    31
  a q/Q to the Sigma, Pair and converse rules.*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
coinductive
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 42798
diff changeset
    33
  domains "lleq(A)" \<subseteq> "llist(A) * llist(A)"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    34
  intros
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    35
    LNil:  "<LNil, LNil> \<in> lleq(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    36
    LCons: "[| a \<in> A; <l,l'> \<in> lleq(A) |] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    37
            ==> <LCons(a,l), LCons(a,l')> \<in> lleq(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    38
  type_intros  llist.intros
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    39
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    40
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    41
(*Lazy list functions; flip is not definitional!*)
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    42
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    43
  lconst   :: "i => i"  where
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    44
  "lconst(a) == lfp(univ(a), %l. LCons(a,l))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    45
41779
a68f503805ed modernized specifications;
wenzelm
parents: 32960
diff changeset
    46
axiomatization flip :: "i => i"
a68f503805ed modernized specifications;
wenzelm
parents: 32960
diff changeset
    47
where
a68f503805ed modernized specifications;
wenzelm
parents: 32960
diff changeset
    48
  flip_LNil:   "flip(LNil) = LNil" and
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    49
  flip_LCons:  "[| x \<in> bool; l \<in> llist(bool) |] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    50
                ==> flip(LCons(x,l)) = LCons(not(x), flip(l))"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    51
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    52
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    53
(*These commands cause classical reasoning to regard the subset relation
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    54
  as primitive, not reducing it to membership*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    55
  
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    56
declare subsetI [rule del]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    57
        subsetCE [rule del]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    58
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    59
declare subset_refl [intro!] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    60
        cons_subsetI [intro!] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    61
        subset_consI [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    62
        Union_least [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    63
        UN_least [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    64
        Un_least [intro!]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
    65
        Inter_greatest [intro!]
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    66
        Int_greatest [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    67
        RepFun_subset [intro!]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
    68
        Un_upper1 [intro!]
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    69
        Un_upper2 [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    70
        Int_lower1 [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    71
        Int_lower2 [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    72
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    73
(*An elimination rule, for type-checking*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    74
inductive_cases LConsE: "LCons(a,l) \<in> llist(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    75
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    76
(*Proving freeness results*)
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 42798
diff changeset
    77
lemma LCons_iff: "LCons(a,l)=LCons(a',l') \<longleftrightarrow> a=a' & l=l'"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    78
by auto
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    79
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    80
lemma LNil_LCons_iff: "~ LNil=LCons(a,l)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    81
by auto
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    82
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    83
(*
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    84
lemma llist_unfold: "llist(A) = {0} <+> (A <*> llist(A))";
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    85
let open llist  val rew = rewrite_rule con_defs in  
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    86
by (fast_tac (claset() addSIs (subsetI ::map rew intros) addEs [rew elim]) 1)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    87
end
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    88
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    89
*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    90
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    91
(*** Lemmas to justify using "llist" in other recursive type definitions ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    92
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    93
lemma llist_mono: "A \<subseteq> B ==> llist(A) \<subseteq> llist(B)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    94
apply (unfold llist.defs )
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    95
apply (rule gfp_mono)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    96
apply (rule llist.bnd_mono)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    97
apply (assumption | rule quniv_mono basic_monos)+
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    98
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
    99
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   100
(** Closure of quniv(A) under llist -- why so complex?  Its a gfp... **)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   101
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   102
declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   103
        QPair_subset_univ [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   104
        empty_subsetI [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   105
        one_in_quniv [THEN qunivD, intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   106
declare qunivD [dest!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   107
declare Ord_in_Ord [elim!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   108
46935
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   109
lemma llist_quniv_lemma:
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   110
     "Ord(i) ==> l \<in> llist(quniv(A)) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> univ(eclose(A))"
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   111
proof (induct i arbitrary: l rule: trans_induct)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   112
  case (step i l)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   113
  show ?case using `l \<in> llist(quniv(A))`
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   114
  proof (cases l rule: llist.cases)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   115
    case LNil thus ?thesis
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   116
      by (simp add: QInl_def QInr_def llist.con_defs)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   117
  next
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   118
    case (LCons a l) thus ?thesis using step.hyps
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   119
    proof (simp add: QInl_def QInr_def llist.con_defs)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   120
      show "<1; <a; l>> \<inter> Vset(i) \<subseteq> univ(eclose(A))" using LCons `Ord(i)`
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   121
        by (fast intro: step Ord_trans Int_lower1 [THEN subset_trans])
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   122
    qed
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   123
  qed
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   124
qed
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   125
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   126
lemma llist_quniv: "llist(quniv(A)) \<subseteq> quniv(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   127
apply (rule qunivI [THEN subsetI])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   128
apply (rule Int_Vset_subset)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   129
apply (assumption | rule llist_quniv_lemma)+
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   130
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   131
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   132
lemmas llist_subset_quniv =
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   133
       subset_trans [OF llist_mono llist_quniv]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   134
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   135
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   136
(*** Lazy List Equality: lleq ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   137
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   138
declare QPair_Int_Vset_subset_UN [THEN subset_trans, intro!] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   139
        QPair_mono [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   140
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   141
declare Ord_in_Ord [elim!] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   142
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   143
(*Lemma for proving finality.  Unfold the lazy list; use induction hypothesis*)
46935
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   144
lemma lleq_Int_Vset_subset:
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   145
     "Ord(i) ==> <l,l'> \<in> lleq(A) \<Longrightarrow> l \<inter> Vset(i) \<subseteq> l'"
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   146
proof (induct i arbitrary: l l' rule: trans_induct)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   147
  case (step i l l')
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   148
  show ?case using `\<langle>l, l'\<rangle> \<in> lleq(A)`
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   149
  proof (cases rule: lleq.cases)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   150
    case LNil thus ?thesis
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   151
      by (auto simp add: QInr_def llist.con_defs)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   152
  next
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   153
    case (LCons a l l') thus ?thesis using step.hyps
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   154
      by (auto simp add: QInr_def llist.con_defs intro: Ord_trans)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   155
  qed
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   156
qed
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   157
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   158
(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   159
lemma lleq_symmetric: "<l,l'> \<in> lleq(A) ==> <l',l> \<in> lleq(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   160
apply (erule lleq.coinduct [OF converseI]) 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12867
diff changeset
   161
apply (rule lleq.dom_subset [THEN converse_type], safe)
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12867
diff changeset
   162
apply (erule lleq.cases, blast+)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   163
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   164
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   165
lemma lleq_implies_equal: "<l,l'> \<in> lleq(A) ==> l=l'"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   166
apply (rule equalityI)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   167
apply (assumption | rule lleq_Int_Vset_subset [THEN Int_Vset_subset] | 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   168
       erule lleq_symmetric)+
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   169
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   170
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   171
lemma equal_llist_implies_leq:
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   172
     "[| l=l';  l \<in> llist(A) |] ==> <l,l'> \<in> lleq(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   173
apply (rule_tac X = "{<l,l>. l \<in> llist (A) }" in lleq.coinduct)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   174
apply blast
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   175
apply safe
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12867
diff changeset
   176
apply (erule_tac a=l in llist.cases, fast+)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   177
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   178
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   179
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   180
(*** Lazy List Functions ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   181
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   182
(*Examples of coinduction for type-checking and to prove llist equations*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   183
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   184
(*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   185
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   186
lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), %l. LCons(a,l))"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   187
apply (unfold llist.con_defs )
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   188
apply (rule bnd_monoI)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   189
apply (blast intro: A_subset_univ QInr_subset_univ)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   190
apply (blast intro: subset_refl QInr_mono QPair_mono)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   191
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   192
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   193
(* lconst(a) = LCons(a,lconst(a)) *)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   194
lemmas lconst = def_lfp_unfold [OF lconst_def lconst_fun_bnd_mono]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   195
lemmas lconst_subset = lconst_def [THEN def_lfp_subset]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   196
lemmas member_subset_Union_eclose = arg_into_eclose [THEN Union_upper]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   197
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   198
lemma lconst_in_quniv: "a \<in> A ==> lconst(a) \<in> quniv(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   199
apply (rule lconst_subset [THEN subset_trans, THEN qunivI])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   200
apply (erule arg_into_eclose [THEN eclose_subset, THEN univ_mono])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   201
done
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   202
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   203
lemma lconst_type: "a \<in> A ==> lconst(a): llist(A)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   204
apply (rule singletonI [THEN llist.coinduct])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   205
apply (erule lconst_in_quniv [THEN singleton_subsetI])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   206
apply (fast intro!: lconst)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   207
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   208
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   209
(*** flip --- equations merely assumed; certain consequences proved ***)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   210
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   211
declare flip_LNil [simp] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   212
        flip_LCons [simp] 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   213
        not_type [simp]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   214
46822
95f1e700b712 mathematical symbols for Isabelle/ZF example theories
paulson
parents: 42798
diff changeset
   215
lemma bool_Int_subset_univ: "b \<in> bool ==> b \<inter> X \<subseteq> univ(eclose(A))"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   216
by (fast intro: Int_lower1 [THEN subset_trans] elim!: boolE)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   217
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   218
declare not_type [intro!]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   219
declare bool_Int_subset_univ [intro]
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   220
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   221
(*Reasoning borrowed from lleq.ML; a similar proof works for all
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   222
  "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
46935
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   223
lemma flip_llist_quniv_lemma:
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   224
     "Ord(i) ==> l \<in> llist(bool) \<Longrightarrow> flip(l) \<inter> Vset(i) \<subseteq> univ(eclose(bool))"
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   225
proof (induct i arbitrary: l rule: trans_induct)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   226
  case (step i l)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   227
  show ?case using `l \<in> llist(bool)`
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   228
  proof (cases l rule: llist.cases)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   229
    case LNil thus ?thesis
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   230
      by (simp, simp add: QInl_def QInr_def llist.con_defs)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   231
  next
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   232
    case (LCons a l) thus ?thesis using step.hyps
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   233
    proof (simp, simp add: QInl_def QInr_def llist.con_defs)
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   234
      show "<1; <not(a); flip(l)>> \<inter> Vset(i) \<subseteq> univ(eclose(bool))" using LCons step.hyps
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   235
        by (auto intro: Ord_trans) 
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   236
    qed
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   237
  qed
38ecb2dc3636 structured case and induct rules
paulson
parents: 46822
diff changeset
   238
qed
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   239
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   240
lemma flip_in_quniv: "l \<in> llist(bool) ==> flip(l) \<in> quniv(bool)"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12867
diff changeset
   241
by (rule flip_llist_quniv_lemma [THEN Int_Vset_subset, THEN qunivI], assumption+)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   242
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   243
lemma flip_type: "l \<in> llist(bool) ==> flip(l): llist(bool)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   244
apply (rule_tac X = "{flip (l) . l \<in> llist (bool) }" in llist.coinduct)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   245
apply blast
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   246
apply (fast intro!: flip_in_quniv)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   247
apply (erule RepFunE)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12867
diff changeset
   248
apply (erule_tac a=la in llist.cases, auto)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   249
done
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   250
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   251
lemma flip_flip: "l \<in> llist(bool) ==> flip(flip(l)) = l"
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   252
apply (rule_tac X1 = "{<flip (flip (l)),l> . l \<in> llist (bool) }" in 
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   253
       lleq.coinduct [THEN lleq_implies_equal])
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   254
apply blast
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   255
apply (fast intro!: flip_type)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   256
apply (erule RepFunE)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 12867
diff changeset
   257
apply (erule_tac a=la in llist.cases)
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   258
apply (simp_all add: flip_type not_not)
5c900a821a7c New-style versions of these old examples
paulson
parents: 11354
diff changeset
   259
done
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   260
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   261
end
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
   262