src/HOL/Datatype_Examples/Stream_Processor.thy
author paulson <lp15@cam.ac.uk>
Wed, 19 Aug 2020 14:58:02 +0100
changeset 72171 7075fe8ffd76
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
more lex fixes
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58309
a09ec6daaa19 renamed 'BNF_Examples' to 'Datatype_Examples' (cf. 'datatypes.pdf')
blanchet
parents: 57641
diff changeset
     1
(*  Title:      HOL/Datatype_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
57634
blanchet
parents: 57398
diff changeset
     6
Stream processors---a syntactic representation of continuous functions on streams.
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     7
*)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
     8
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
     9
section \<open>Stream Processors---A Syntactic Representation of Continuous Functions on Streams\<close>
54961
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
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63167
diff changeset
    12
imports "HOL-Library.Stream" "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
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 60754
diff changeset
    15
declare [[typedef_overloaded]]
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 60754
diff changeset
    16
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 60754
diff changeset
    17
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
    18
section \<open>Continuous Functions on Streams\<close>
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    19
62694
f50d7efc8fe3 added two 'corec' examples
blanchet
parents: 61260
diff changeset
    20
datatype ('a, 'b, 'c) sp\<^sub>\<mu> =
f50d7efc8fe3 added two 'corec' examples
blanchet
parents: 61260
diff changeset
    21
  Get "'a \<Rightarrow> ('a, 'b, 'c) sp\<^sub>\<mu>"
f50d7efc8fe3 added two 'corec' examples
blanchet
parents: 61260
diff changeset
    22
| Put "'b" "'c"
f50d7efc8fe3 added two 'corec' examples
blanchet
parents: 61260
diff changeset
    23
f50d7efc8fe3 added two 'corec' examples
blanchet
parents: 61260
diff changeset
    24
codatatype ('a, 'b) sp\<^sub>\<nu> =
f50d7efc8fe3 added two 'corec' examples
blanchet
parents: 61260
diff changeset
    25
  In (out: "('a, 'b, ('a, 'b) sp\<^sub>\<nu>) sp\<^sub>\<mu>")
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    26
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
    27
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
    28
  "run\<^sub>\<mu> (Get f) s = run\<^sub>\<mu> (f (shd s)) (stl s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    29
| "run\<^sub>\<mu> (Put b sp) s = ((b, sp), s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    30
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    31
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
    32
  "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
    33
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    34
primcorec copy :: "('a, 'a) sp\<^sub>\<nu>" where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    35
  "copy = In (Get (\<lambda>a. Put a copy))"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    36
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    37
lemma run\<^sub>\<nu>_copy: "run\<^sub>\<nu> copy s = s"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    38
  by (coinduction arbitrary: s) simp
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    39
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
    40
text \<open>
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    41
To use the function package for the definition of composition the
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    42
wellfoundedness of the subtree relation needs to be proved first.
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
    43
\<close>
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    44
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    45
definition "sub \<equiv> {(f a, Get f) | a f. True}"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    46
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    47
lemma subI[intro]: "(f a, Get f) \<in> sub"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    48
  unfolding sub_def by blast
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    49
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    50
lemma wf_sub[simp, intro]: "wf sub"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    51
proof (rule wfUNIVI)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    52
  fix P  :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> bool" and x
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    53
  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
    54
  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
    55
  show "P x" by (induct x) (auto intro: I)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    56
qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    57
72171
7075fe8ffd76 more lex fixes
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    58
lemma neq_Get [simp]: "f b \<noteq> Get f"
7075fe8ffd76 more lex fixes
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    59
  by (metis subI wf_asym wf_sub)
7075fe8ffd76 more lex fixes
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    60
7075fe8ffd76 more lex fixes
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    61
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    62
function
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    63
  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
    64
  (infixl "o\<^sub>\<mu>" 65)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    65
where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    66
  "Put b sp o\<^sub>\<mu> fsp = Put b (sp, In fsp)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    67
| "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
    68
| "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
    69
by pat_completeness auto
72171
7075fe8ffd76 more lex fixes
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    70
termination
7075fe8ffd76 more lex fixes
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    71
  by (relation "lex_prod sub sub") auto
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    72
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    73
primcorec sp\<^sub>\<nu>_comp (infixl "o\<^sub>\<nu>" 65) where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    74
  "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
    75
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    76
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
    77
proof (rule ext, unfold comp_apply)
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    78
  fix s
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    79
  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
    80
  proof (coinduction arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    81
    case Eq_stream
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    82
    show ?case
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    83
    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
    84
      case (1 b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    85
      show ?case by (auto simp add: 1[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    86
    next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    87
      case (2 f b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    88
      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
    89
    next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    90
      case (3 f h)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    91
      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
    92
    qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    93
  qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    94
qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    95
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
    96
text \<open>Alternative definition of composition using primrec instead of function\<close>
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    97
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
    98
primrec sp\<^sub>\<mu>_comp2R  where
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
    99
  "sp\<^sub>\<mu>_comp2R f (Put b sp) = f b (out sp)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   100
| "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
   101
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
   102
primrec sp\<^sub>\<mu>_comp2 (infixl "o\<^sup>*\<^sub>\<mu>" 65) where
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   103
  "Put b sp o\<^sup>*\<^sub>\<mu> fsp = Put b (sp, In fsp)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   104
| "Get f o\<^sup>*\<^sub>\<mu> fsp = sp\<^sub>\<mu>_comp2R ((o\<^sup>*\<^sub>\<mu>) o f) fsp"
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   105
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   106
primcorec sp\<^sub>\<nu>_comp2 (infixl "o\<^sup>*\<^sub>\<nu>" 65) where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   107
  "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
   108
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   109
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
   110
proof (rule ext, unfold comp_apply)
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   111
  fix s
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   112
  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
   113
  proof (coinduction arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   114
    case Eq_stream
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   115
    show ?case
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   116
    proof (induct "out sp" arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   117
      case (Put b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   118
      show ?case by (auto simp add: Put[symmetric])
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   119
    next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   120
      case (Get f)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   121
      then show ?case
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   122
      proof (induct "out sp'" arbitrary: sp sp' s)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   123
        case (Put b sp'')
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   124
        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
   125
      next
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   126
        case (Get h)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   127
        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
   128
      qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   129
    qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   130
  qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   131
qed
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   132
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
   133
text \<open>The two definitions are equivalent\<close>
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   134
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   135
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
   136
  by (induct sp sp' rule: sp\<^sub>\<mu>_comp.induct) auto
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   137
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   138
(*will be provided by the package*)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   139
lemma sp\<^sub>\<mu>_rel_map_map[unfolded vimage2p_def, simp]:
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   140
  "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
   141
  rel_sp\<^sub>\<mu> (BNF_Def.vimage2p f1 g1 R1) (BNF_Def.vimage2p f2 g2 R2) sp sp'"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
   142
by (tactic \<open>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   143
  let
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   144
    val ks = 1 upto 2;
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   145
    val ctxt = \<^context>;
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   146
  in
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   147
    BNF_Tactics.unfold_thms_tac ctxt
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   148
      @{thms sp\<^sub>\<mu>.rel_compp sp\<^sub>\<mu>.rel_conversep sp\<^sub>\<mu>.rel_Grp vimage2p_Grp} THEN
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   149
    HEADGOAL (EVERY' [resolve_tac ctxt [iffI],
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   150
      resolve_tac ctxt @{thms relcomppI},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   151
      resolve_tac ctxt @{thms GrpI},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   152
      resolve_tac ctxt [refl],
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   153
      resolve_tac ctxt [CollectI],
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   154
      BNF_Util.CONJ_WRAP' (K (resolve_tac ctxt @{thms subset_UNIV})) ks,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   155
      resolve_tac ctxt @{thms relcomppI}, assume_tac ctxt,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   156
      resolve_tac ctxt @{thms conversepI},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   157
      resolve_tac ctxt @{thms GrpI},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   158
      resolve_tac ctxt [refl],
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   159
      resolve_tac ctxt [CollectI],
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   160
      BNF_Util.CONJ_WRAP' (K (resolve_tac ctxt @{thms subset_UNIV})) ks,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   161
      REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE conversepE GrpE},
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   162
      hyp_subst_tac ctxt,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59498
diff changeset
   163
      assume_tac ctxt])
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   164
  end
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
   165
\<close>)
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   166
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   167
lemma sp\<^sub>\<mu>_rel_self: "\<lbrakk>(=) \<le> R1; (=) \<le> R2\<rbrakk> \<Longrightarrow> rel_sp\<^sub>\<mu> R1 R2 x x"
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   168
  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
   169
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   170
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
   171
  by (coinduction arbitrary: sp sp') (auto intro!: sp\<^sub>\<mu>_rel_self)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   172
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   173
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62694
diff changeset
   174
section \<open>Generalization to an Arbitrary BNF as Codomain\<close>
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   175
57206
d9be905d6283 changed syntax of map: and rel: arguments to BNF-based datatypes
blanchet
parents: 56942
diff changeset
   176
bnf_axiomatization ('a, 'b) F for map: F
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   177
57641
dc59f147b27d more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents: 57634
diff changeset
   178
notation BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
55091
c43394c2e5ec compile
blanchet
parents: 55076
diff changeset
   179
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   180
definition \<theta> :: "('p,'a) F * 'b \<Rightarrow> ('p,'a * 'b) F" where
57641
dc59f147b27d more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
wenzelm
parents: 57634
diff changeset
   181
  "\<theta> xb = F id \<langle>id, \<lambda> a. (snd xb)\<rangle> (fst xb)"
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   182
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   183
(* The strength laws for \<theta>: *)
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55804
diff changeset
   184
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
   185
  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
   186
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   187
definition assl :: "'a * ('b * 'c) \<Rightarrow> ('a * 'b) * 'c" where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   188
  "assl abc = ((fst abc, fst (snd abc)), snd (snd abc))"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   189
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   190
lemma \<theta>_rid: "F id fst o \<theta> = fst"
55066
blanchet
parents: 54961
diff changeset
   191
  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
   192
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55804
diff changeset
   193
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
   194
  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
   195
58310
91ea607a34d8 updated news
blanchet
parents: 58309
diff changeset
   196
datatype ('a, 'b, 'c) spF\<^sub>\<mu> = GetF "'a \<Rightarrow> ('a, 'b, 'c) spF\<^sub>\<mu>" | PutF "('b,'c) F"
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   197
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
   198
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   199
codatatype 'b JF = Ctor (dtor: "('b, 'b JF) F")
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   200
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   201
(* Definition of run for an arbitrary final coalgebra as codomain: *)
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   202
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
   203
primrec
54961
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   204
  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
   205
where
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   206
  "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
   207
| "runF\<^sub>\<mu> (PutF x) s = (x, s)"
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   208
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   209
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
   210
  "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
   211
e60428f432bc new codatatype example: stream processors
traytel
parents:
diff changeset
   212
end