src/HOL/BNF_Examples/Stream_Processor.thy
author haftmann
Sat, 05 Jul 2014 11:01:53 +0200
changeset 57514 bdc2c6b40bf2
parent 57398 882091eb1e9a
child 57634 efc00b9b8680
permissions -rw-r--r--
prefer ac_simps collections over separate name bindings for add and mult
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55075
b3d0a02a756d dissolved BNF session
blanchet
parents: 55071
diff changeset
     1
(*  Title:      HOL/BNF_Examples/Stream_Processor.thy
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     3
    Author:     Andrei Popescu, TU Muenchen
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     4
    Copyright   2014
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     5
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     6
Stream processors---a syntactic representation of continuous functions on streams
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     7
*)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     8
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     9
header {* Stream Processors *}
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    10
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    11
theory Stream_Processor
56942
5fff4dc31d34 bnf_decl -> bnf_axiomatization
traytel
parents: 55932
diff changeset
    12
imports Stream "~~/src/HOL/Library/BNF_Axiomatization"
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    13
begin
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    14
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    15
section {* Continuous Functions on Streams *}
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    16
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    17
datatype_new ('a, 'b, 'c) sp\<^sub>\<mu> = Get "'a \<Rightarrow> ('a, 'b, 'c) sp\<^sub>\<mu>" | Put "'b" "'c"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    18
codatatype ('a, 'b) sp\<^sub>\<nu> = In (out: "('a, 'b, ('a, 'b) sp\<^sub>\<nu>) sp\<^sub>\<mu>")
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    19
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55091
diff changeset
    20
primrec run\<^sub>\<mu> :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> 'a stream \<Rightarrow> ('b \<times> 'c) \<times> 'a stream" where
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    21
  "run\<^sub>\<mu> (Get f) s = run\<^sub>\<mu> (f (shd s)) (stl s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    22
| "run\<^sub>\<mu> (Put b sp) s = ((b, sp), s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    23
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    24
primcorec run\<^sub>\<nu> :: "('a, 'b) sp\<^sub>\<nu> \<Rightarrow> 'a stream \<Rightarrow> 'b stream" where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    25
  "run\<^sub>\<nu> sp s = (let ((h, sp), s) = run\<^sub>\<mu> (out sp) s in h ## run\<^sub>\<nu> sp s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    26
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    27
primcorec copy :: "('a, 'a) sp\<^sub>\<nu>" where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    28
  "copy = In (Get (\<lambda>a. Put a copy))"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    29
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    30
lemma run\<^sub>\<nu>_copy: "run\<^sub>\<nu> copy s = s"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    31
  by (coinduction arbitrary: s) simp
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    32
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    33
text {*
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    34
To use the function package for the definition of composition the
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    35
wellfoundedness of the subtree relation needs to be proved first.
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    36
*}
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    37
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    38
definition "sub \<equiv> {(f a, Get f) | a f. True}"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    39
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    40
lemma subI[intro]: "(f a, Get f) \<in> sub"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    41
  unfolding sub_def by blast
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    42
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    43
lemma wf_sub[simp, intro]: "wf sub"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    44
proof (rule wfUNIVI)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    45
  fix P  :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> bool" and x
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    46
  assume "\<forall>x. (\<forall>y. (y, x) \<in> sub \<longrightarrow> P y) \<longrightarrow> P x"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    47
  hence I: "\<And>x. (\<forall>y. (\<exists>a f. y = f a \<and> x = Get f) \<longrightarrow> P y) \<Longrightarrow> P x" unfolding sub_def by blast
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    48
  show "P x" by (induct x) (auto intro: I)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    49
qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    50
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    51
function
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    52
  sp\<^sub>\<mu>_comp :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> ('d, 'a, ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu> \<Rightarrow> ('d, 'b, 'c \<times> ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu>"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    53
  (infixl "o\<^sub>\<mu>" 65)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    54
where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    55
  "Put b sp o\<^sub>\<mu> fsp = Put b (sp, In fsp)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    56
| "Get f o\<^sub>\<mu> Put b sp = f b o\<^sub>\<mu> out sp"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    57
| "Get f o\<^sub>\<mu> Get g = Get (\<lambda>a. Get f o\<^sub>\<mu> g a)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    58
by pat_completeness auto
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    59
termination by (relation "lex_prod sub sub") auto
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    60
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    61
primcorec sp\<^sub>\<nu>_comp (infixl "o\<^sub>\<nu>" 65) where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    62
  "out (sp o\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sub>\<nu> sp') (out sp o\<^sub>\<mu> out sp')"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    63
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    64
lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp: "run\<^sub>\<nu> (sp o\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
55066
blanchet
parents: 54961
diff changeset
    65
proof (rule ext, unfold comp_apply)
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    66
  fix s
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    67
  show "run\<^sub>\<nu> (sp o\<^sub>\<nu> sp') s = run\<^sub>\<nu> sp (run\<^sub>\<nu> sp' s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    68
  proof (coinduction arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    69
    case Eq_stream
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    70
    show ?case
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    71
    proof (induct "out sp" "out sp'" arbitrary: sp sp' s rule: sp\<^sub>\<mu>_comp.induct)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    72
      case (1 b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    73
      show ?case by (auto simp add: 1[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    74
    next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    75
      case (2 f b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    76
      from 2(1)[of "In (f b)" sp''] show ?case by (simp add: 2(2,3)[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    77
    next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    78
      case (3 f h)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    79
      from 3(1)[of _ "shd s" "In (h (shd s))", OF 3(2)] show ?case by (simp add: 3(2,3)[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    80
    qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    81
  qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    82
qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    83
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55091
diff changeset
    84
text {* Alternative definition of composition using primrec instead of function *}
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    85
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55091
diff changeset
    86
primrec sp\<^sub>\<mu>_comp2R  where
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    87
  "sp\<^sub>\<mu>_comp2R f (Put b sp) = f b (out sp)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    88
| "sp\<^sub>\<mu>_comp2R f (Get h) = Get (sp\<^sub>\<mu>_comp2R f o h)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    89
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55091
diff changeset
    90
primrec sp\<^sub>\<mu>_comp2 (infixl "o\<^sup>*\<^sub>\<mu>" 65) where
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    91
  "Put b sp o\<^sup>*\<^sub>\<mu> fsp = Put b (sp, In fsp)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    92
| "Get f o\<^sup>*\<^sub>\<mu> fsp = sp\<^sub>\<mu>_comp2R (op o\<^sup>*\<^sub>\<mu> o f) fsp"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    93
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    94
primcorec sp\<^sub>\<nu>_comp2 (infixl "o\<^sup>*\<^sub>\<nu>" 65) where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    95
  "out (sp o\<^sup>*\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sup>*\<^sub>\<nu> sp') (out sp o\<^sup>*\<^sub>\<mu> out sp')"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    96
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    97
lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp2: "run\<^sub>\<nu> (sp o\<^sup>*\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
55066
blanchet
parents: 54961
diff changeset
    98
proof (rule ext, unfold comp_apply)
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    99
  fix s
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   100
  show "run\<^sub>\<nu> (sp o\<^sup>*\<^sub>\<nu> sp') s = run\<^sub>\<nu> sp (run\<^sub>\<nu> sp' s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   101
  proof (coinduction arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   102
    case Eq_stream
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   103
    show ?case
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   104
    proof (induct "out sp" arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   105
      case (Put b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   106
      show ?case by (auto simp add: Put[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   107
    next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   108
      case (Get f)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   109
      then show ?case
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   110
      proof (induct "out sp'" arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   111
        case (Put b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   112
        from Put(2)[of "In (f b)" sp''] show ?case by (simp add: Put(1,3)[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   113
      next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   114
        case (Get h)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   115
        from Get(1)[OF _ Get(3,4), of "In (h (shd s))"] show ?case by (simp add: Get(2,4)[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   116
      qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   117
    qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   118
  qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   119
qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   120
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   121
text {* The two definitions are equivalent *}
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   122
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   123
lemma sp\<^sub>\<mu>_comp_sp\<^sub>\<mu>_comp2[simp]: "sp o\<^sub>\<mu> sp' = sp o\<^sup>*\<^sub>\<mu> sp'"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   124
  by (induct sp sp' rule: sp\<^sub>\<mu>_comp.induct) auto
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   125
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   126
(*will be provided by the package*)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   127
lemma sp\<^sub>\<mu>_rel_map_map[unfolded vimage2p_def, simp]:
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   128
  "rel_sp\<^sub>\<mu> R1 R2 (map_sp\<^sub>\<mu> f1 f2 sp) (map_sp\<^sub>\<mu> g1 g2 sp') =
57398
882091eb1e9a merged two small theory files
blanchet
parents: 57206
diff changeset
   129
  rel_sp\<^sub>\<mu> (BNF_Def.vimage2p f1 g1 R1) (BNF_Def.vimage2p f2 g2 R2) sp sp'"
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   130
by (tactic {*
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   131
  let val ks = 1 upto 2;
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   132
  in
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   133
    BNF_Tactics.unfold_thms_tac @{context}
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   134
      @{thms sp\<^sub>\<mu>.rel_compp sp\<^sub>\<mu>.rel_conversep sp\<^sub>\<mu>.rel_Grp vimage2p_Grp} THEN
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   135
    HEADGOAL (EVERY' [rtac iffI, rtac @{thm relcomppI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   136
      BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, rtac @{thm relcomppI}, atac,
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   137
      rtac @{thm conversepI}, rtac @{thm GrpI}, rtac refl, rtac CollectI,
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   138
      BNF_Util.CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks,
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   139
      REPEAT_DETERM o eresolve_tac @{thms relcomppE conversepE GrpE},
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   140
      hyp_subst_tac @{context}, atac])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   141
  end
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   142
*})
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   143
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   144
lemma sp\<^sub>\<mu>_rel_self: "\<lbrakk>op = \<le> R1; op = \<le> R2\<rbrakk> \<Longrightarrow> rel_sp\<^sub>\<mu> R1 R2 x x"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   145
  by (erule (1) predicate2D[OF sp\<^sub>\<mu>.rel_mono]) (simp only: sp\<^sub>\<mu>.rel_eq)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   146
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   147
lemma sp\<^sub>\<nu>_comp_sp\<^sub>\<nu>_comp2: "sp o\<^sub>\<nu> sp' = sp o\<^sup>*\<^sub>\<nu> sp'"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   148
  by (coinduction arbitrary: sp sp') (auto intro!: sp\<^sub>\<mu>_rel_self)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   149
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   150
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   151
section {* Generalization to an Arbitrary BNF as Codomain *}
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   152
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 56942
diff changeset
   153
bnf_axiomatization ('a, 'b) F for map: F
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   154
55091
c43394c2e5ec compile
blanchet
parents: 55076
diff changeset
   155
notation BNF_Def.convol ("<_ , _>")
c43394c2e5ec compile
blanchet
parents: 55076
diff changeset
   156
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   157
definition \<theta> :: "('p,'a) F * 'b \<Rightarrow> ('p,'a * 'b) F" where
55091
c43394c2e5ec compile
blanchet
parents: 55076
diff changeset
   158
  "\<theta> xb = F id <id, \<lambda> a. (snd xb)> (fst xb)"
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   159
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   160
(* The strength laws for \<theta>: *)
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55804
diff changeset
   161
lemma \<theta>_natural: "F id (map_prod f g) o \<theta> = \<theta> o map_prod (F id f) g"
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55804
diff changeset
   162
  unfolding \<theta>_def F.map_comp comp_def id_apply convol_def map_prod_def split_beta fst_conv snd_conv ..
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   163
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   164
definition assl :: "'a * ('b * 'c) \<Rightarrow> ('a * 'b) * 'c" where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   165
  "assl abc = ((fst abc, fst (snd abc)), snd (snd abc))"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   166
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   167
lemma \<theta>_rid: "F id fst o \<theta> = fst"
55066
blanchet
parents: 54961
diff changeset
   168
  unfolding \<theta>_def F.map_comp F.map_id comp_def id_apply convol_def fst_conv sym[OF id_def] ..
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   169
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55804
diff changeset
   170
lemma \<theta>_assl: "F id assl o \<theta> = \<theta> o map_prod \<theta> id o assl"
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55804
diff changeset
   171
  unfolding assl_def \<theta>_def F.map_comp comp_def id_apply convol_def map_prod_def split fst_conv snd_conv ..
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   172
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   173
datatype_new ('a, 'b, 'c) spF\<^sub>\<mu> = GetF "'a \<Rightarrow> ('a, 'b, 'c) spF\<^sub>\<mu>" | PutF "('b,'c) F"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   174
codatatype ('a, 'b) spF\<^sub>\<nu> = InF (outF: "('a, 'b, ('a, 'b) spF\<^sub>\<nu>) spF\<^sub>\<mu>")
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   175
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   176
codatatype 'b JF = Ctor (dtor: "('b, 'b JF) F")
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   177
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   178
(* Definition of run for an arbitrary final coalgebra as codomain: *)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   179
55530
3dfb724db099 renamed 'primrec_new' to 'primrec', overriding the old command (which it still uses as a fallback for old-style datatypes)
blanchet
parents: 55091
diff changeset
   180
primrec
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   181
  runF\<^sub>\<mu> :: "('a, 'b, ('a, 'b) spF\<^sub>\<nu>) spF\<^sub>\<mu> \<Rightarrow> 'a stream \<Rightarrow> (('b, ('a, 'b) spF\<^sub>\<nu>) F) \<times> 'a stream" 
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   182
where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   183
  "runF\<^sub>\<mu> (GetF f) s = (runF\<^sub>\<mu> o f) (shd s) (stl s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   184
| "runF\<^sub>\<mu> (PutF x) s = (x, s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   185
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   186
primcorec runF\<^sub>\<nu> :: "('a, 'b) spF\<^sub>\<nu> \<Rightarrow> 'a stream \<Rightarrow> 'b JF" where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   187
  "runF\<^sub>\<nu> sp s = (let (x, s) = runF\<^sub>\<mu> (outF sp) s in Ctor (F id (\<lambda> sp. runF\<^sub>\<nu> sp s) x))"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   188
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   189
end