src/HOLCF/Streams.thy
author oheimb
Mon, 12 Apr 2004 12:18:48 +0200
changeset 14535 7cb26928e70d
child 14981 e73f8140af78
permissions -rw-r--r--
added Streams.thy (with stream concatenation etc.)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14535
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     1
(*  Title: 	HOLCF/Streams.thy
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     2
    ID:         $Id$
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     3
    Author: 	Borislav Gajanovic and David von Oheimb
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     5
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     6
Stream domains with concatenation.
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     7
TODO: HOLCF/ex/Stream.* should be integrated into this file.
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     8
*)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
     9
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    10
theory Streams = Stream :
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    11
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    12
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    13
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    14
lemma stream_neq_UU: "x~=UU ==> EX a as. x=a&&as & a~=UU"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    15
by (simp add: stream_exhaust_eq,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    16
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    17
lemma stream_prefix1: "[| x<<y; xs<<ys |] ==> x&&xs << y&&ys"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    18
by (insert stream_prefix' [of y "x&&xs" ys],force)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    19
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    20
lemma stream_take_le_mono : "k<=n ==> stream_take k$s1 << stream_take n$s1"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    21
apply (insert chain_stream_take [of s1])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    22
by (drule chain_mono3,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    23
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    24
lemma mono_stream_take: "s1 << s2 ==> stream_take n$s1 << stream_take n$s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    25
by (simp add: monofun_cfun_arg)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    26
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    27
lemma stream_take_prefix [simp]: "stream_take n$s << s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    28
apply (subgoal_tac "s=(LUB n. stream_take n$s)")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    29
 apply (erule ssubst, rule is_ub_thelub)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    30
 apply (simp only: chain_stream_take)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    31
by (simp only: stream_reach2)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    32
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    33
lemma stream_take_take_less:"stream_take k$(stream_take n$s) << stream_take k$s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    34
by (rule monofun_cfun_arg,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    35
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    36
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    37
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    38
lemma slen_rt_mono: "#s2 <= #s1 ==> #(rt$s2) <= #(rt$s1)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    39
apply (rule stream.casedist [of s1])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    40
 apply (rule stream.casedist [of s2],simp+)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    41
by (rule stream.casedist [of s2],auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    42
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    43
lemma slen_take_lemma4 [rule_format]: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    44
  "!s. stream_take n$s ~= s --> #(stream_take n$s) = Fin n"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    45
apply (induct_tac n,auto simp add: Fin_0)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    46
apply (case_tac "s=UU",simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    47
by (drule stream_neq_UU,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    48
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    49
lemma slen_take_lemma5: "#(stream_take n$s) <= Fin n"; 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    50
apply (case_tac "stream_take n$s = s")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    51
 apply (simp add: slen_take_eq_rev)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    52
by (simp add: slen_take_lemma4)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    53
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    54
lemma stream_take_idempotent [simp]: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    55
 "stream_take n$(stream_take n$s) = stream_take n$s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    56
apply (case_tac "stream_take n$s = s")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    57
apply (auto,insert slen_take_lemma4 [of n s]);
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    58
by (auto,insert slen_take_lemma1 [of "stream_take n$s" n],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    59
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    60
lemma stream_take_take_Suc [simp]: "stream_take n$(stream_take (Suc n)$s) = 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    61
                                    stream_take n$s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    62
apply (simp add: po_eq_conv,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    63
 apply (simp add: stream_take_take_less)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    64
apply (subgoal_tac "stream_take n$s = stream_take n$(stream_take n$s)")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    65
 apply (erule ssubst)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    66
 apply (rule_tac monofun_cfun_arg)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    67
 apply (insert chain_stream_take [of s])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    68
by (simp add: chain_def,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    69
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    70
lemma mono_stream_take_pred: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    71
  "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    72
                       stream_take n$s1 << stream_take n$s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    73
by (drule mono_stream_take [of _ _ n],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    74
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    75
lemma stream_take_lemma10 [rule_format]:
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    76
  "ALL k<=n. stream_take n$s1 << stream_take n$s2 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    77
                             --> stream_take k$s1 << stream_take k$s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    78
apply (induct_tac n,simp,clarsimp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    79
apply (case_tac "k=Suc n",blast)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    80
apply (erule_tac x="k" in allE)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    81
by (drule mono_stream_take_pred,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    82
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    83
lemma stream_take_finite [simp]: "stream_finite (stream_take n$s)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    84
apply (simp add: stream.finite_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    85
by (rule_tac x="n" in exI,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    86
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    87
lemma slen_stream_take_finite [simp]: "#(stream_take n$s) ~= \<infinity>"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    88
by (simp add: slen_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    89
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    90
lemma stream_take_Suc_neq: "stream_take (Suc n)$s ~=s ==> 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    91
                     stream_take n$s ~= stream_take (Suc n)$s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    92
apply auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    93
apply (subgoal_tac "stream_take n$s ~=s")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    94
 apply (insert slen_take_lemma4 [of n s],auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    95
apply (rule stream.casedist [of s],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    96
apply (simp add: inat_defs split:inat_splits)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    97
by (simp add: slen_take_lemma4)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    98
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
    99
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   100
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   101
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   102
consts
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   103
 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   104
  i_rt :: "nat => 'a stream => 'a stream" (* chops the first i elements *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   105
  i_th :: "nat => 'a stream => 'a"        (* the i-th element *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   106
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   107
  sconc         :: "'a stream => 'a stream => 'a stream" (infixr "ooo" 65) 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   108
  constr_sconc  :: "'a stream => 'a stream => 'a stream" (* constructive *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   109
  constr_sconc' :: "nat => 'a stream => 'a stream => 'a stream" 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   110
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   111
defs
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   112
  i_rt_def: "i_rt == %i s. iterate i rt s"  
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   113
  i_th_def: "i_th == %i s. ft$(i_rt i s)" 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   114
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   115
  sconc_def: "s1 ooo s2 == case #s1 of 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   116
                       Fin n => (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   117
                     | \<infinity>     => s1" 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   118
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   119
  constr_sconc_def: "constr_sconc s1 s2 == case #s1 of 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   120
                                             Fin n => constr_sconc' n s1 s2 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   121
                                           | \<infinity>    => s1"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   122
primrec 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   123
  constr_sconc'_0:   "constr_sconc' 0 s1 s2 = s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   124
  constr_sconc'_Suc: "constr_sconc' (Suc n) s1 s2 = ft$s1 && 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   125
                                                    constr_sconc' n (rt$s1) s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   126
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   127
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   128
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   129
   section "i_rt"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   130
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   131
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   132
lemma i_rt_UU [simp]: "i_rt n UU = UU"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   133
apply (simp add: i_rt_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   134
by (rule iterate.induct,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   135
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   136
lemma i_rt_0 [simp]: "i_rt 0 s = s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   137
by (simp add: i_rt_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   138
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   139
lemma i_rt_Suc [simp]: "a ~= UU ==> i_rt (Suc n) (a&&s) = i_rt n s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   140
by (simp add: i_rt_def iterate_Suc2 del: iterate_Suc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   141
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   142
lemma i_rt_Suc_forw: "i_rt (Suc n) s = i_rt n (rt$s)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   143
by (simp only: i_rt_def iterate_Suc2)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   144
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   145
lemma i_rt_Suc_back:"i_rt (Suc n) s = rt$(i_rt n s)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   146
by (simp only: i_rt_def,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   147
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   148
lemma i_rt_mono: "x << s ==> i_rt n x  << i_rt n s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   149
by (simp add: i_rt_def monofun_rt_mult)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   150
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   151
lemma i_rt_ij_lemma: "Fin (i + j) <= #x ==> Fin j <= #(i_rt i x)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   152
by (simp add: i_rt_def slen_rt_mult)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   153
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   154
lemma slen_i_rt_mono: "#s2 <= #s1 ==> #(i_rt n s2) <= #(i_rt n s1)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   155
apply (induct_tac n,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   156
apply (simp add: i_rt_Suc_back)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   157
by (drule slen_rt_mono,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   158
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   159
lemma i_rt_take_lemma1 [rule_format]: "ALL s. i_rt n (stream_take n$s) = UU"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   160
apply (induct_tac n); 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   161
 apply (simp add: i_rt_Suc_back,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   162
apply (case_tac "s=UU",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   163
by (drule stream_neq_UU,simp add: i_rt_Suc_forw,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   164
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   165
lemma i_rt_slen: "(i_rt n s = UU) = (stream_take n$s = s)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   166
apply auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   167
 apply (insert i_rt_ij_lemma [of n "Suc 0" s]);
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   168
 apply (subgoal_tac "#(i_rt n s)=0")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   169
  apply (case_tac "stream_take n$s = s",simp+)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   170
  apply (insert slen_take_eq [of n s],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   171
  apply (simp add: inat_defs split:inat_splits)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   172
 apply (simp add: slen_take_eq )
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   173
by (simp, insert i_rt_take_lemma1 [of n s],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   174
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   175
lemma i_rt_lemma_slen: "#s=Fin n ==> i_rt n s = UU"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   176
by (simp add: i_rt_slen slen_take_lemma1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   177
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   178
lemma stream_finite_i_rt [simp]: "stream_finite (i_rt n s) = stream_finite s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   179
apply (induct_tac n, auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   180
 apply (rule stream.casedist [of "s"], auto simp del: i_rt_Suc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   181
by (simp add: i_rt_Suc_back stream_finite_rt_eq)+
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   182
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   183
lemma take_i_rt_len_lemma: "ALL sl x j t. Fin sl = #x & n <= sl &
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   184
                            #(stream_take n$x) = Fin t & #(i_rt n x)= Fin j 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   185
                                              --> Fin (j + t) = #x"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   186
apply (induct_tac n,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   187
 apply (simp add: inat_defs)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   188
apply (case_tac "x=UU",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   189
 apply (simp add: inat_defs)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   190
apply (drule stream_neq_UU,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   191
apply (subgoal_tac "EX k. Fin k = #as",clarify)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   192
 apply (erule_tac x="k" in allE)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   193
 apply (erule_tac x="as" in allE,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   194
 apply (erule_tac x="THE p. Suc p = t" in allE,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   195
   apply (simp add: inat_defs split:inat_splits)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   196
  apply (simp add: inat_defs split:inat_splits)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   197
  apply (simp only: the_equality)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   198
 apply (simp add: inat_defs split:inat_splits)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   199
 apply force
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   200
by (simp add: inat_defs split:inat_splits)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   201
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   202
lemma take_i_rt_len: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   203
"[| Fin sl = #x; n <= sl; #(stream_take n$x) = Fin t; #(i_rt n x) = Fin j |] ==>
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   204
    Fin (j + t) = #x"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   205
by (blast intro: take_i_rt_len_lemma [rule_format])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   206
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   207
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   208
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   209
   section "i_th"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   210
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   211
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   212
lemma i_th_i_rt_step:
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   213
"[| i_th n s1 << i_th n s2; i_rt (Suc n) s1 << i_rt (Suc n) s2 |] ==> 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   214
   i_rt n s1 << i_rt n s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   215
apply (simp add: i_th_def i_rt_Suc_back)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   216
apply (rule stream.casedist [of "i_rt n s1"],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   217
apply (rule stream.casedist [of "i_rt n s2"],auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   218
by (drule stream_prefix1,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   219
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   220
lemma i_th_stream_take_Suc [rule_format]: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   221
 "ALL s. i_th n (stream_take (Suc n)$s) = i_th n s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   222
apply (induct_tac n,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   223
 apply (simp add: i_th_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   224
 apply (case_tac "s=UU",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   225
 apply (drule stream_neq_UU,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   226
apply (case_tac "s=UU",simp add: i_th_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   227
apply (drule stream_neq_UU,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   228
by (simp add: i_th_def i_rt_Suc_forw)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   229
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   230
lemma last_lemma10: "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==> 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   231
                     i_th n s1 << i_th n s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   232
apply (rule i_th_stream_take_Suc [THEN subst])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   233
apply (rule i_th_stream_take_Suc [THEN subst]) back
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   234
apply (simp add: i_th_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   235
apply (rule monofun_cfun_arg)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   236
by (erule i_rt_mono)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   237
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   238
lemma i_th_last: "i_th n s && UU = i_rt n (stream_take (Suc n)$s)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   239
apply (insert surjectiv_scons [of "i_rt n (stream_take (Suc n)$s)"])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   240
apply (rule i_th_stream_take_Suc [THEN subst])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   241
apply (simp add: i_th_def  i_rt_Suc_back [symmetric])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   242
by (simp add: i_rt_take_lemma1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   243
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   244
lemma i_th_last_eq: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   245
"i_th n s1 = i_th n s2 ==> i_rt n (stream_take (Suc n)$s1) = i_rt n (stream_take (Suc n)$s2)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   246
apply (insert i_th_last [of n s1])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   247
apply (insert i_th_last [of n s2])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   248
by auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   249
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   250
lemma i_th_prefix_lemma:
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   251
"[| k <= n; stream_take (Suc n)$s1 << stream_take (Suc n)$s2 |] ==> 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   252
    i_th k s1 << i_th k s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   253
apply (subgoal_tac "stream_take (Suc k)$s1 << stream_take (Suc k)$s2")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   254
 apply (simp add: last_lemma10)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   255
by (blast intro: stream_take_lemma10)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   256
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   257
lemma take_i_rt_prefix_lemma1: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   258
  "stream_take (Suc n)$s1 << stream_take (Suc n)$s2 ==>
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   259
   i_rt (Suc n) s1 << i_rt (Suc n) s2 ==> 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   260
   i_rt n s1 << i_rt n s2 & stream_take n$s1 << stream_take n$s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   261
apply auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   262
 apply (insert i_th_prefix_lemma [of n n s1 s2])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   263
 apply (rule i_th_i_rt_step,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   264
by (drule mono_stream_take_pred,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   265
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   266
lemma take_i_rt_prefix_lemma: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   267
"[| stream_take n$s1 << stream_take n$s2; i_rt n s1 << i_rt n s2 |] ==> s1 << s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   268
apply (case_tac "n=0",simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   269
apply (insert neq0_conv [of n])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   270
apply (insert not0_implies_Suc [of n],auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   271
apply (subgoal_tac "stream_take 0$s1 << stream_take 0$s2 & 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   272
                    i_rt 0 s1 << i_rt 0 s2")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   273
 defer 1
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   274
 apply (rule zero_induct,blast)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   275
 apply (blast dest: take_i_rt_prefix_lemma1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   276
by simp
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   277
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   278
lemma streams_prefix_lemma: "(s1 << s2) = 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   279
  (stream_take n$s1 << stream_take n$s2 & i_rt n s1 << i_rt n s2)"; 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   280
apply auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   281
  apply (simp add: monofun_cfun_arg)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   282
 apply (simp add: i_rt_mono)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   283
by (erule take_i_rt_prefix_lemma,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   284
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   285
lemma streams_prefix_lemma1:
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   286
 "[| stream_take n$s1 = stream_take n$s2; i_rt n s1 = i_rt n s2 |] ==> s1 = s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   287
apply (simp add: po_eq_conv,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   288
 apply (insert streams_prefix_lemma)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   289
 by blast+
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   290
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   291
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   292
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   293
   section "sconc"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   294
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   295
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   296
lemma UU_sconc [simp]: " UU ooo s = s "
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   297
by (simp add: sconc_def inat_defs)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   298
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   299
lemma scons_neq_UU: "a~=UU ==> a && s ~=UU"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   300
by auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   301
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   302
lemma singleton_sconc [rule_format, simp]: "x~=UU --> (x && UU) ooo y = x && y"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   303
apply (simp add: sconc_def inat_defs split:inat_splits,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   304
apply (rule someI2_ex,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   305
 apply (rule_tac x="x && y" in exI,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   306
apply (simp add: i_rt_Suc_forw)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   307
apply (case_tac "xa=UU",simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   308
by (drule stream_neq_UU,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   309
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   310
lemma ex_sconc [rule_format]: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   311
  "ALL k y. #x = Fin k --> (EX w. stream_take k$w = x & i_rt k w = y)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   312
apply (case_tac "#x")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   313
 apply (rule stream_finite_ind [of x],auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   314
  apply (simp add: stream.finite_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   315
  apply (drule slen_take_lemma1,blast)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   316
 apply (simp add: inat_defs split:inat_splits)+
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   317
apply (erule_tac x="y" in allE,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   318
by (rule_tac x="a && w" in exI,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   319
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   320
lemma rt_sconc1: "Fin n = #x ==> i_rt n (x ooo y) = y"; 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   321
apply (simp add: sconc_def inat_defs split:inat_splits , arith?,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   322
apply (rule someI2_ex,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   323
by (drule ex_sconc,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   324
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   325
lemma sconc_inj2: "\<lbrakk>Fin n = #x; x ooo y = x ooo z\<rbrakk> \<Longrightarrow> y = z"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   326
apply (frule_tac y=y in rt_sconc1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   327
by (auto elim: rt_sconc1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   328
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   329
lemma sconc_UU [simp]:"s ooo UU = s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   330
apply (case_tac "#s")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   331
 apply (simp add: sconc_def inat_defs)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   332
 apply (rule someI2_ex)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   333
  apply (rule_tac x="s" in exI)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   334
  apply auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   335
   apply (drule slen_take_lemma1,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   336
  apply (simp add: i_rt_lemma_slen)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   337
 apply (drule slen_take_lemma1,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   338
 apply (simp add: i_rt_slen)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   339
by (simp add: sconc_def inat_defs)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   340
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   341
lemma stream_take_sconc [simp]: "Fin n = #x ==> stream_take n$(x ooo y) = x"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   342
apply (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   343
apply (simp add: inat_defs split:inat_splits,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   344
apply (rule someI2_ex,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   345
by (drule ex_sconc,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   346
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   347
lemma scons_sconc [rule_format,simp]: "a~=UU --> (a && x) ooo y = a && x ooo y"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   348
apply (case_tac "#x",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   349
 apply (simp add: sconc_def) 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   350
 apply (rule someI2_ex)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   351
  apply (drule ex_sconc,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   352
 apply (rule someI2_ex,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   353
  apply (simp add: i_rt_Suc_forw)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   354
  apply (rule_tac x="a && x" in exI,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   355
 apply (case_tac "xa=UU",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   356
  apply (drule_tac s="stream_take nat$x" in scons_neq_UU)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   357
  apply (simp add: i_rt_Suc_forw)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   358
 apply (drule stream_neq_UU,clarsimp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   359
 apply (drule streams_prefix_lemma1,simp+)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   360
by (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   361
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   362
lemma ft_sconc: "x ~= UU ==> ft$(x ooo y) = ft$x"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   363
by (rule stream.casedist [of x],auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   364
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   365
lemma sconc_assoc: "(x ooo y) ooo z = x ooo y ooo z"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   366
apply (case_tac "#x")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   367
 apply (rule stream_finite_ind [of x],auto simp del: scons_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   368
  apply (simp add: stream.finite_def del: scons_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   369
  apply (drule slen_take_lemma1,auto simp del: scons_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   370
 apply (case_tac "a = UU", auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   371
by (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   372
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   373
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   374
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   375
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   376
lemma sconc_mono: "y << y' ==> x ooo y << x ooo y'"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   377
apply (case_tac "#x")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   378
 apply (rule stream_finite_ind [of "x"])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   379
   apply (auto simp add: stream.finite_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   380
  apply (drule slen_take_lemma1,blast)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   381
 by (simp add: stream_prefix',auto simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   382
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   383
lemma sconc_mono1 [simp]: "x << x ooo y"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   384
by (rule sconc_mono [of UU, simplified])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   385
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   386
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   387
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   388
lemma empty_sconc [simp]: "(x ooo y = UU) = (x = UU & y = UU)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   389
apply (case_tac "#x",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   390
   by (insert sconc_mono1 [of x y],auto);
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   391
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   392
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   393
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   394
lemma rt_sconc [rule_format, simp]: "s~=UU --> rt$(s ooo x) = rt$s ooo x"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   395
by (rule stream.casedist,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   396
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   397
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   398
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   399
lemma sconc_lemma [rule_format, simp]: "ALL s. stream_take n$s ooo i_rt n s = s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   400
apply (induct_tac n,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   401
apply (case_tac "s=UU",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   402
by (drule stream_neq_UU,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   403
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   404
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   405
   subsection "pointwise equality"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   406
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   407
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   408
lemma ex_last_stream_take_scons: "stream_take (Suc n)$s = 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   409
                     stream_take n$s ooo i_rt n (stream_take (Suc n)$s)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   410
by (insert sconc_lemma [of n "stream_take (Suc n)$s"],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   411
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   412
lemma i_th_stream_take_eq: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   413
"!!n. ALL n. i_th n s1 = i_th n s2 ==> stream_take n$s1 = stream_take n$s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   414
apply (induct_tac n,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   415
apply (subgoal_tac "stream_take (Suc na)$s1 =
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   416
                    stream_take na$s1 ooo i_rt na (stream_take (Suc na)$s1)")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   417
 apply (subgoal_tac "i_rt na (stream_take (Suc na)$s1) = 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   418
                    i_rt na (stream_take (Suc na)$s2)")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   419
  apply (subgoal_tac "stream_take (Suc na)$s2 = 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   420
                    stream_take na$s2 ooo i_rt na (stream_take (Suc na)$s2)")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   421
   apply (insert ex_last_stream_take_scons,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   422
  apply blast
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   423
 apply (erule_tac x="na" in allE)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   424
 apply (insert i_th_last_eq [of _ s1 s2])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   425
by blast+
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   426
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   427
lemma pointwise_eq_lemma[rule_format]: "ALL n. i_th n s1 = i_th n s2 ==> s1 = s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   428
by (insert i_th_stream_take_eq [THEN stream.take_lemmas],blast)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   429
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   430
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   431
   subsection "finiteness"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   432
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   433
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   434
lemma slen_sconc_finite1:
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   435
  "[| #(x ooo y) = Infty; Fin n = #x |] ==> #y = Infty"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   436
apply (case_tac "#y ~= Infty",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   437
apply (simp only: slen_infinite [symmetric])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   438
apply (drule_tac y=y in rt_sconc1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   439
apply (insert stream_finite_i_rt [of n "x ooo y"])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   440
by (simp add: slen_infinite)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   441
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   442
lemma slen_sconc_infinite1: "#x=Infty ==> #(x ooo y) = Infty"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   443
by (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   444
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   445
lemma slen_sconc_infinite2: "#y=Infty ==> #(x ooo y) = Infty"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   446
apply (case_tac "#x")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   447
 apply (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   448
 apply (rule someI2_ex)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   449
  apply (drule ex_sconc,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   450
 apply (erule contrapos_pp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   451
 apply (insert stream_finite_i_rt)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   452
 apply (simp add: slen_infinite ,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   453
by (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   454
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   455
lemma sconc_finite: "(#x~=Infty & #y~=Infty) = (#(x ooo y)~=Infty)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   456
apply auto
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   457
  apply (case_tac "#x",auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   458
  apply (erule contrapos_pp,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   459
  apply (erule slen_sconc_finite1,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   460
 apply (drule slen_sconc_infinite1 [of _ y],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   461
by (drule slen_sconc_infinite2 [of _ x],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   462
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   463
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   464
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   465
lemma slen_sconc_mono3: "[| Fin n = #x; Fin k = #(x ooo y) |] ==> n <= k"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   466
apply (insert slen_mono [of "x" "x ooo y"])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   467
by (simp add: inat_defs split: inat_splits)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   468
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   469
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   470
   subsection "finite slen"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   471
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   472
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   473
lemma slen_sconc: "[| Fin n = #x; Fin m = #y |] ==> #(x ooo y) = Fin (n + m)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   474
apply (case_tac "#(x ooo y)")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   475
 apply (frule_tac y=y in rt_sconc1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   476
 apply (insert take_i_rt_len [of "THE j. Fin j = #(x ooo y)" "x ooo y" n n m],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   477
 apply (insert slen_sconc_mono3 [of n x _ y],simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   478
by (insert sconc_finite [of x y],auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   479
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   480
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   481
   subsection "flat prefix"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   482
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   483
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   484
lemma sconc_prefix: "(s1::'a::flat stream) << s2 ==> EX t. s1 ooo t = s2"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   485
apply (case_tac "#s1")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   486
 apply (subgoal_tac "stream_take nat$s1 = stream_take nat$s2");
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   487
  apply (rule_tac x="i_rt nat s2" in exI)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   488
  apply (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   489
  apply (rule someI2_ex)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   490
   apply (drule ex_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   491
   apply (simp,clarsimp,drule streams_prefix_lemma1)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   492
   apply (simp+,rule slen_take_lemma3 [rule_format, of _ s1 s2]);
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   493
  apply (simp+,rule_tac x="UU" in exI)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   494
apply (insert slen_take_lemma3 [rule_format, of _ s1 s2]);
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   495
by (rule stream.take_lemmas,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   496
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   497
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   498
   subsection "continuity"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   499
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   500
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   501
lemma chain_sconc: "chain S ==> chain (%i. (x ooo S i))"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   502
by (simp add: chain_def,auto simp add: sconc_mono)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   503
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   504
lemma chain_scons: "chain S ==> chain (%i. a && S i)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   505
apply (simp add: chain_def,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   506
by (rule monofun_cfun_arg,simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   507
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   508
lemma contlub_scons: "contlub (%x. a && x)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   509
by (simp add: contlub_Rep_CFun2)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   510
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   511
lemma contlub_scons_lemma: "chain S ==> (LUB i. a && S i) = a && (LUB i. S i)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   512
apply (insert contlub_scons [of a])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   513
by (simp only: contlub)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   514
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   515
lemma finite_lub_sconc: "chain Y ==> (stream_finite x) ==> 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   516
                        (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   517
apply (rule stream_finite_ind [of x])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   518
 apply (auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   519
apply (subgoal_tac "(LUB i. a && (s ooo Y i)) = a && (LUB i. s ooo Y i)")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   520
 by (force,blast dest: contlub_scons_lemma chain_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   521
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   522
lemma contlub_sconc_lemma: 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   523
  "chain Y ==> (LUB i. x ooo Y i) = (x ooo (LUB i. Y i))"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   524
apply (case_tac "#x=Infty")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   525
 apply (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   526
 prefer 2
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   527
 apply (drule finite_lub_sconc,auto simp add: slen_infinite)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   528
apply (simp add: slen_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   529
apply (insert lub_const [of x] unique_lub [of _ x _])
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   530
by (auto simp add: lub)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   531
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   532
lemma contlub_sconc: "contlub (%y. x ooo y)"; 
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   533
by (rule contlubI, insert contlub_sconc_lemma [of _ x], simp);
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   534
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   535
lemma monofun_sconc: "monofun (%y. x ooo y)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   536
by (simp add: monofun sconc_mono)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   537
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   538
lemma cont_sconc: "cont (%y. x ooo y)"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   539
apply (rule  monocontlub2cont)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   540
 apply (rule monofunI, simp add: sconc_mono)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   541
by (rule contlub_sconc);
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   542
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   543
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   544
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   545
   section "constr_sconc"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   546
(* ----------------------------------------------------------------------- *)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   547
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   548
lemma constr_sconc_UUs [simp]: "constr_sconc UU s = s"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   549
by (simp add: constr_sconc_def inat_defs)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   550
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   551
lemma "x ooo y = constr_sconc x y"
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   552
apply (case_tac "#x")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   553
 apply (rule stream_finite_ind [of x],auto simp del: scons_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   554
  defer 1
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   555
  apply (simp add: constr_sconc_def del: scons_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   556
  apply (case_tac "#s")
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   557
   apply (simp add: inat_defs)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   558
   apply (case_tac "a=UU",auto simp del: scons_sconc)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   559
   apply (simp)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   560
  apply (simp add: sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   561
 apply (simp add: constr_sconc_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   562
apply (simp add: stream.finite_def)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   563
by (drule slen_take_lemma1,auto)
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   564
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents:
diff changeset
   565
end