src/HOL/HOLCF/Map_Functions.thy
author wenzelm
Mon, 29 May 2017 19:34:07 +0200
changeset 65959 47309113ee4d
parent 65380 ae93953746fc
child 67312 0d25e02759b7
permissions -rw-r--r--
clarified view column; tuned whitespace;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41297
diff changeset
     1
(*  Title:      HOL/HOLCF/Map_Functions.thy
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
     3
*)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
     5
section \<open>Map functions for various types\<close>
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
     6
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
     7
theory Map_Functions
65380
ae93953746fc eliminated Plain_HOLCF.thy (see also 8e92772bc0e8): it was modeled after HOL/Plain.thy which was discontinued later;
wenzelm
parents: 62175
diff changeset
     8
imports Deflation Sprod Ssum Sfun Up
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
     9
begin
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    10
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
    11
subsection \<open>Map operator for continuous function space\<close>
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    12
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    13
default_sort cpo
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    14
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    15
definition
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    16
  cfun_map :: "('b \<rightarrow> 'a) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> ('a \<rightarrow> 'c) \<rightarrow> ('b \<rightarrow> 'd)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    17
where
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    18
  "cfun_map = (\<Lambda> a b f x. b\<cdot>(f\<cdot>(a\<cdot>x)))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    19
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    20
lemma cfun_map_beta [simp]: "cfun_map\<cdot>a\<cdot>b\<cdot>f\<cdot>x = b\<cdot>(f\<cdot>(a\<cdot>x))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    21
unfolding cfun_map_def by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    22
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    23
lemma cfun_map_ID: "cfun_map\<cdot>ID\<cdot>ID = ID"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    24
unfolding cfun_eq_iff by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    25
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    26
lemma cfun_map_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    27
  "cfun_map\<cdot>f1\<cdot>g1\<cdot>(cfun_map\<cdot>f2\<cdot>g2\<cdot>p) =
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    28
    cfun_map\<cdot>(\<Lambda> x. f2\<cdot>(f1\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    29
by (rule cfun_eqI) simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    30
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    31
lemma ep_pair_cfun_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    32
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    33
  shows "ep_pair (cfun_map\<cdot>p1\<cdot>e2) (cfun_map\<cdot>e1\<cdot>p2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    34
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    35
  interpret e1p1: ep_pair e1 p1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    36
  interpret e2p2: ep_pair e2 p2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    37
  fix f show "cfun_map\<cdot>e1\<cdot>p2\<cdot>(cfun_map\<cdot>p1\<cdot>e2\<cdot>f) = f"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    38
    by (simp add: cfun_eq_iff)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    39
  fix g show "cfun_map\<cdot>p1\<cdot>e2\<cdot>(cfun_map\<cdot>e1\<cdot>p2\<cdot>g) \<sqsubseteq> g"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    40
    apply (rule cfun_belowI, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    41
    apply (rule below_trans [OF e2p2.e_p_below])
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    42
    apply (rule monofun_cfun_arg)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    43
    apply (rule e1p1.e_p_below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    44
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    45
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    46
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    47
lemma deflation_cfun_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    48
  assumes "deflation d1" and "deflation d2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    49
  shows "deflation (cfun_map\<cdot>d1\<cdot>d2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    50
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    51
  interpret d1: deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    52
  interpret d2: deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    53
  fix f
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    54
  show "cfun_map\<cdot>d1\<cdot>d2\<cdot>(cfun_map\<cdot>d1\<cdot>d2\<cdot>f) = cfun_map\<cdot>d1\<cdot>d2\<cdot>f"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    55
    by (simp add: cfun_eq_iff d1.idem d2.idem)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    56
  show "cfun_map\<cdot>d1\<cdot>d2\<cdot>f \<sqsubseteq> f"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    57
    apply (rule cfun_belowI, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    58
    apply (rule below_trans [OF d2.below])
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    59
    apply (rule monofun_cfun_arg)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    60
    apply (rule d1.below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    61
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    62
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    63
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    64
lemma finite_range_cfun_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    65
  assumes a: "finite (range (\<lambda>x. a\<cdot>x))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    66
  assumes b: "finite (range (\<lambda>y. b\<cdot>y))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    67
  shows "finite (range (\<lambda>f. cfun_map\<cdot>a\<cdot>b\<cdot>f))"  (is "finite (range ?h)")
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    68
proof (rule finite_imageD)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    69
  let ?f = "\<lambda>g. range (\<lambda>x. (a\<cdot>x, g\<cdot>x))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    70
  show "finite (?f ` range ?h)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    71
  proof (rule finite_subset)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    72
    let ?B = "Pow (range (\<lambda>x. a\<cdot>x) \<times> range (\<lambda>y. b\<cdot>y))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    73
    show "?f ` range ?h \<subseteq> ?B"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    74
      by clarsimp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    75
    show "finite ?B"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    76
      by (simp add: a b)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    77
  qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    78
  show "inj_on ?f (range ?h)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    79
  proof (rule inj_onI, rule cfun_eqI, clarsimp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    80
    fix x f g
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    81
    assume "range (\<lambda>x. (a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x)))) = range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    82
    hence "range (\<lambda>x. (a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x)))) \<subseteq> range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    83
      by (rule equalityD1)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    84
    hence "(a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x))) \<in> range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    85
      by (simp add: subset_eq)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    86
    then obtain y where "(a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x))) = (a\<cdot>y, b\<cdot>(g\<cdot>(a\<cdot>y)))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    87
      by (rule rangeE)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    88
    thus "b\<cdot>(f\<cdot>(a\<cdot>x)) = b\<cdot>(g\<cdot>(a\<cdot>x))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    89
      by clarsimp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    90
  qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    91
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    92
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    93
lemma finite_deflation_cfun_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    94
  assumes "finite_deflation d1" and "finite_deflation d2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    95
  shows "finite_deflation (cfun_map\<cdot>d1\<cdot>d2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    96
proof (rule finite_deflation_intro)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    97
  interpret d1: finite_deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    98
  interpret d2: finite_deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
    99
  have "deflation d1" and "deflation d2" by fact+
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   100
  thus "deflation (cfun_map\<cdot>d1\<cdot>d2)" by (rule deflation_cfun_map)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   101
  have "finite (range (\<lambda>f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   102
    using d1.finite_range d2.finite_range
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   103
    by (rule finite_range_cfun_map)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   104
  thus "finite {f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f = f}"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   105
    by (rule finite_range_imp_finite_fixes)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   106
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   107
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   108
text \<open>Finite deflations are compact elements of the function space\<close>
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   109
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   110
lemma finite_deflation_imp_compact: "finite_deflation d \<Longrightarrow> compact d"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   111
apply (frule finite_deflation_imp_deflation)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   112
apply (subgoal_tac "compact (cfun_map\<cdot>d\<cdot>d\<cdot>d)")
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   113
apply (simp add: cfun_map_def deflation.idem eta_cfun)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   114
apply (rule finite_deflation.compact)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   115
apply (simp only: finite_deflation_cfun_map)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   116
done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   117
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   118
subsection \<open>Map operator for product type\<close>
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   119
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   120
definition
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   121
  prod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<times> 'c \<rightarrow> 'b \<times> 'd"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   122
where
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   123
  "prod_map = (\<Lambda> f g p. (f\<cdot>(fst p), g\<cdot>(snd p)))"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   124
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   125
lemma prod_map_Pair [simp]: "prod_map\<cdot>f\<cdot>g\<cdot>(x, y) = (f\<cdot>x, g\<cdot>y)"
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   126
unfolding prod_map_def by simp
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   127
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   128
lemma prod_map_ID: "prod_map\<cdot>ID\<cdot>ID = ID"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   129
unfolding cfun_eq_iff by auto
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   130
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   131
lemma prod_map_map:
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   132
  "prod_map\<cdot>f1\<cdot>g1\<cdot>(prod_map\<cdot>f2\<cdot>g2\<cdot>p) =
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   133
    prod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   134
by (induct p) simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   135
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   136
lemma ep_pair_prod_map:
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   137
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   138
  shows "ep_pair (prod_map\<cdot>e1\<cdot>e2) (prod_map\<cdot>p1\<cdot>p2)"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   139
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   140
  interpret e1p1: ep_pair e1 p1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   141
  interpret e2p2: ep_pair e2 p2 by fact
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   142
  fix x show "prod_map\<cdot>p1\<cdot>p2\<cdot>(prod_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   143
    by (induct x) simp
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   144
  fix y show "prod_map\<cdot>e1\<cdot>e2\<cdot>(prod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   145
    by (induct y) (simp add: e1p1.e_p_below e2p2.e_p_below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   146
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   147
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   148
lemma deflation_prod_map:
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   149
  assumes "deflation d1" and "deflation d2"
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   150
  shows "deflation (prod_map\<cdot>d1\<cdot>d2)"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   151
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   152
  interpret d1: deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   153
  interpret d2: deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   154
  fix x
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   155
  show "prod_map\<cdot>d1\<cdot>d2\<cdot>(prod_map\<cdot>d1\<cdot>d2\<cdot>x) = prod_map\<cdot>d1\<cdot>d2\<cdot>x"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   156
    by (induct x) (simp add: d1.idem d2.idem)
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   157
  show "prod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   158
    by (induct x) (simp add: d1.below d2.below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   159
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   160
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   161
lemma finite_deflation_prod_map:
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   162
  assumes "finite_deflation d1" and "finite_deflation d2"
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   163
  shows "finite_deflation (prod_map\<cdot>d1\<cdot>d2)"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   164
proof (rule finite_deflation_intro)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   165
  interpret d1: finite_deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   166
  interpret d2: finite_deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   167
  have "deflation d1" and "deflation d2" by fact+
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   168
  thus "deflation (prod_map\<cdot>d1\<cdot>d2)" by (rule deflation_prod_map)
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   169
  have "{p. prod_map\<cdot>d1\<cdot>d2\<cdot>p = p} \<subseteq> {x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   170
    by clarsimp
41297
01b2de947cff rename function cprod_map to prod_map
huffman
parents: 41291
diff changeset
   171
  thus "finite {p. prod_map\<cdot>d1\<cdot>d2\<cdot>p = p}"
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   172
    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   173
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   174
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   175
subsection \<open>Map function for lifted cpo\<close>
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   176
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   177
definition
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   178
  u_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a u \<rightarrow> 'b u"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   179
where
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   180
  "u_map = (\<Lambda> f. fup\<cdot>(up oo f))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   181
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   182
lemma u_map_strict [simp]: "u_map\<cdot>f\<cdot>\<bottom> = \<bottom>"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   183
unfolding u_map_def by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   184
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   185
lemma u_map_up [simp]: "u_map\<cdot>f\<cdot>(up\<cdot>x) = up\<cdot>(f\<cdot>x)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   186
unfolding u_map_def by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   187
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   188
lemma u_map_ID: "u_map\<cdot>ID = ID"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   189
unfolding u_map_def by (simp add: cfun_eq_iff eta_cfun)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   190
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   191
lemma u_map_map: "u_map\<cdot>f\<cdot>(u_map\<cdot>g\<cdot>p) = u_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>p"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   192
by (induct p) simp_all
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   193
41291
752d81c2ce25 add lemma u_map_oo
huffman
parents: 40774
diff changeset
   194
lemma u_map_oo: "u_map\<cdot>(f oo g) = u_map\<cdot>f oo u_map\<cdot>g"
752d81c2ce25 add lemma u_map_oo
huffman
parents: 40774
diff changeset
   195
by (simp add: cfcomp1 u_map_map eta_cfun)
752d81c2ce25 add lemma u_map_oo
huffman
parents: 40774
diff changeset
   196
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   197
lemma ep_pair_u_map: "ep_pair e p \<Longrightarrow> ep_pair (u_map\<cdot>e) (u_map\<cdot>p)"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 58880
diff changeset
   198
apply standard
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   199
apply (case_tac x, simp, simp add: ep_pair.e_inverse)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   200
apply (case_tac y, simp, simp add: ep_pair.e_p_below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   201
done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   202
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   203
lemma deflation_u_map: "deflation d \<Longrightarrow> deflation (u_map\<cdot>d)"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 58880
diff changeset
   204
apply standard
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   205
apply (case_tac x, simp, simp add: deflation.idem)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   206
apply (case_tac x, simp, simp add: deflation.below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   207
done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   208
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   209
lemma finite_deflation_u_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   210
  assumes "finite_deflation d" shows "finite_deflation (u_map\<cdot>d)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   211
proof (rule finite_deflation_intro)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   212
  interpret d: finite_deflation d by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   213
  have "deflation d" by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   214
  thus "deflation (u_map\<cdot>d)" by (rule deflation_u_map)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   215
  have "{x. u_map\<cdot>d\<cdot>x = x} \<subseteq> insert \<bottom> ((\<lambda>x. up\<cdot>x) ` {x. d\<cdot>x = x})"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   216
    by (rule subsetI, case_tac x, simp_all)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   217
  thus "finite {x. u_map\<cdot>d\<cdot>x = x}"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   218
    by (rule finite_subset, simp add: d.finite_fixes)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   219
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   220
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   221
subsection \<open>Map function for strict products\<close>
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   222
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   223
default_sort pcpo
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   224
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   225
definition
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   226
  sprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<otimes> 'c \<rightarrow> 'b \<otimes> 'd"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   227
where
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   228
  "sprod_map = (\<Lambda> f g. ssplit\<cdot>(\<Lambda> x y. (:f\<cdot>x, g\<cdot>y:)))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   229
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   230
lemma sprod_map_strict [simp]: "sprod_map\<cdot>a\<cdot>b\<cdot>\<bottom> = \<bottom>"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   231
unfolding sprod_map_def by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   232
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   233
lemma sprod_map_spair [simp]:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   234
  "x \<noteq> \<bottom> \<Longrightarrow> y \<noteq> \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   235
by (simp add: sprod_map_def)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   236
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   237
lemma sprod_map_spair':
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   238
  "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   239
by (cases "x = \<bottom> \<or> y = \<bottom>") auto
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   240
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   241
lemma sprod_map_ID: "sprod_map\<cdot>ID\<cdot>ID = ID"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   242
unfolding sprod_map_def by (simp add: cfun_eq_iff eta_cfun)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   243
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   244
lemma sprod_map_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   245
  "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   246
    sprod_map\<cdot>f1\<cdot>g1\<cdot>(sprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   247
     sprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   248
apply (induct p, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   249
apply (case_tac "f2\<cdot>x = \<bottom>", simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   250
apply (case_tac "g2\<cdot>y = \<bottom>", simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   251
apply simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   252
done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   253
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   254
lemma ep_pair_sprod_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   255
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   256
  shows "ep_pair (sprod_map\<cdot>e1\<cdot>e2) (sprod_map\<cdot>p1\<cdot>p2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   257
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   258
  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   259
  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   260
  fix x show "sprod_map\<cdot>p1\<cdot>p2\<cdot>(sprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   261
    by (induct x) simp_all
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   262
  fix y show "sprod_map\<cdot>e1\<cdot>e2\<cdot>(sprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   263
    apply (induct y, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   264
    apply (case_tac "p1\<cdot>x = \<bottom>", simp, case_tac "p2\<cdot>y = \<bottom>", simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   265
    apply (simp add: monofun_cfun e1p1.e_p_below e2p2.e_p_below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   266
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   267
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   268
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   269
lemma deflation_sprod_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   270
  assumes "deflation d1" and "deflation d2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   271
  shows "deflation (sprod_map\<cdot>d1\<cdot>d2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   272
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   273
  interpret d1: deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   274
  interpret d2: deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   275
  fix x
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   276
  show "sprod_map\<cdot>d1\<cdot>d2\<cdot>(sprod_map\<cdot>d1\<cdot>d2\<cdot>x) = sprod_map\<cdot>d1\<cdot>d2\<cdot>x"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   277
    apply (induct x, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   278
    apply (case_tac "d1\<cdot>x = \<bottom>", simp, case_tac "d2\<cdot>y = \<bottom>", simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   279
    apply (simp add: d1.idem d2.idem)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   280
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   281
  show "sprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   282
    apply (induct x, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   283
    apply (simp add: monofun_cfun d1.below d2.below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   284
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   285
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   286
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   287
lemma finite_deflation_sprod_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   288
  assumes "finite_deflation d1" and "finite_deflation d2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   289
  shows "finite_deflation (sprod_map\<cdot>d1\<cdot>d2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   290
proof (rule finite_deflation_intro)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   291
  interpret d1: finite_deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   292
  interpret d2: finite_deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   293
  have "deflation d1" and "deflation d2" by fact+
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   294
  thus "deflation (sprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_sprod_map)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   295
  have "{x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq> insert \<bottom>
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   296
        ((\<lambda>(x, y). (:x, y:)) ` ({x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   297
    by (rule subsetI, case_tac x, auto simp add: spair_eq_iff)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   298
  thus "finite {x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   299
    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   300
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   301
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   302
subsection \<open>Map function for strict sums\<close>
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   303
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   304
definition
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   305
  ssum_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<oplus> 'c \<rightarrow> 'b \<oplus> 'd"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   306
where
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   307
  "ssum_map = (\<Lambda> f g. sscase\<cdot>(sinl oo f)\<cdot>(sinr oo g))"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   308
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   309
lemma ssum_map_strict [simp]: "ssum_map\<cdot>f\<cdot>g\<cdot>\<bottom> = \<bottom>"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   310
unfolding ssum_map_def by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   311
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   312
lemma ssum_map_sinl [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   313
unfolding ssum_map_def by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   314
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   315
lemma ssum_map_sinr [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   316
unfolding ssum_map_def by simp
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   317
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   318
lemma ssum_map_sinl': "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   319
by (cases "x = \<bottom>") simp_all
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   320
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   321
lemma ssum_map_sinr': "g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   322
by (cases "x = \<bottom>") simp_all
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   323
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   324
lemma ssum_map_ID: "ssum_map\<cdot>ID\<cdot>ID = ID"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   325
unfolding ssum_map_def by (simp add: cfun_eq_iff eta_cfun)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   326
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   327
lemma ssum_map_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   328
  "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   329
    ssum_map\<cdot>f1\<cdot>g1\<cdot>(ssum_map\<cdot>f2\<cdot>g2\<cdot>p) =
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   330
     ssum_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   331
apply (induct p, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   332
apply (case_tac "f2\<cdot>x = \<bottom>", simp, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   333
apply (case_tac "g2\<cdot>y = \<bottom>", simp, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   334
done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   335
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   336
lemma ep_pair_ssum_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   337
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   338
  shows "ep_pair (ssum_map\<cdot>e1\<cdot>e2) (ssum_map\<cdot>p1\<cdot>p2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   339
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   340
  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   341
  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   342
  fix x show "ssum_map\<cdot>p1\<cdot>p2\<cdot>(ssum_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   343
    by (induct x) simp_all
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   344
  fix y show "ssum_map\<cdot>e1\<cdot>e2\<cdot>(ssum_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   345
    apply (induct y, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   346
    apply (case_tac "p1\<cdot>x = \<bottom>", simp, simp add: e1p1.e_p_below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   347
    apply (case_tac "p2\<cdot>y = \<bottom>", simp, simp add: e2p2.e_p_below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   348
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   349
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   350
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   351
lemma deflation_ssum_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   352
  assumes "deflation d1" and "deflation d2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   353
  shows "deflation (ssum_map\<cdot>d1\<cdot>d2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   354
proof
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   355
  interpret d1: deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   356
  interpret d2: deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   357
  fix x
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   358
  show "ssum_map\<cdot>d1\<cdot>d2\<cdot>(ssum_map\<cdot>d1\<cdot>d2\<cdot>x) = ssum_map\<cdot>d1\<cdot>d2\<cdot>x"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   359
    apply (induct x, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   360
    apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.idem)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   361
    apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.idem)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   362
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   363
  show "ssum_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   364
    apply (induct x, simp)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   365
    apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   366
    apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.below)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   367
    done
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   368
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   369
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   370
lemma finite_deflation_ssum_map:
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   371
  assumes "finite_deflation d1" and "finite_deflation d2"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   372
  shows "finite_deflation (ssum_map\<cdot>d1\<cdot>d2)"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   373
proof (rule finite_deflation_intro)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   374
  interpret d1: finite_deflation d1 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   375
  interpret d2: finite_deflation d2 by fact
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   376
  have "deflation d1" and "deflation d2" by fact+
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   377
  thus "deflation (ssum_map\<cdot>d1\<cdot>d2)" by (rule deflation_ssum_map)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   378
  have "{x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq>
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   379
        (\<lambda>x. sinl\<cdot>x) ` {x. d1\<cdot>x = x} \<union>
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   380
        (\<lambda>x. sinr\<cdot>x) ` {x. d2\<cdot>x = x} \<union> {\<bottom>}"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   381
    by (rule subsetI, case_tac x, simp_all)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   382
  thus "finite {x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   383
    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   384
qed
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   385
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   386
subsection \<open>Map operator for strict function space\<close>
40592
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   387
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   388
definition
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   389
  sfun_map :: "('b \<rightarrow> 'a) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> ('a \<rightarrow>! 'c) \<rightarrow> ('b \<rightarrow>! 'd)"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   390
where
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   391
  "sfun_map = (\<Lambda> a b. sfun_abs oo cfun_map\<cdot>a\<cdot>b oo sfun_rep)"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   392
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   393
lemma sfun_map_ID: "sfun_map\<cdot>ID\<cdot>ID = ID"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   394
  unfolding sfun_map_def
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   395
  by (simp add: cfun_map_ID cfun_eq_iff)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   396
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   397
lemma sfun_map_map:
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   398
  assumes "f2\<cdot>\<bottom> = \<bottom>" and "g2\<cdot>\<bottom> = \<bottom>" shows
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   399
  "sfun_map\<cdot>f1\<cdot>g1\<cdot>(sfun_map\<cdot>f2\<cdot>g2\<cdot>p) =
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   400
    sfun_map\<cdot>(\<Lambda> x. f2\<cdot>(f1\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   401
unfolding sfun_map_def
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   402
by (simp add: cfun_eq_iff strictify_cancel assms cfun_map_map)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   403
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   404
lemma ep_pair_sfun_map:
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   405
  assumes 1: "ep_pair e1 p1"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   406
  assumes 2: "ep_pair e2 p2"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   407
  shows "ep_pair (sfun_map\<cdot>p1\<cdot>e2) (sfun_map\<cdot>e1\<cdot>p2)"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   408
proof
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   409
  interpret e1p1: pcpo_ep_pair e1 p1
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   410
    unfolding pcpo_ep_pair_def by fact
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   411
  interpret e2p2: pcpo_ep_pair e2 p2
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   412
    unfolding pcpo_ep_pair_def by fact
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   413
  fix f show "sfun_map\<cdot>e1\<cdot>p2\<cdot>(sfun_map\<cdot>p1\<cdot>e2\<cdot>f) = f"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   414
    unfolding sfun_map_def
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   415
    apply (simp add: sfun_eq_iff strictify_cancel)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   416
    apply (rule ep_pair.e_inverse)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   417
    apply (rule ep_pair_cfun_map [OF 1 2])
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   418
    done
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   419
  fix g show "sfun_map\<cdot>p1\<cdot>e2\<cdot>(sfun_map\<cdot>e1\<cdot>p2\<cdot>g) \<sqsubseteq> g"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   420
    unfolding sfun_map_def
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   421
    apply (simp add: sfun_below_iff strictify_cancel)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   422
    apply (rule ep_pair.e_p_below)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   423
    apply (rule ep_pair_cfun_map [OF 1 2])
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   424
    done
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   425
qed
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   426
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   427
lemma deflation_sfun_map:
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   428
  assumes 1: "deflation d1"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   429
  assumes 2: "deflation d2"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   430
  shows "deflation (sfun_map\<cdot>d1\<cdot>d2)"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   431
apply (simp add: sfun_map_def)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   432
apply (rule deflation.intro)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   433
apply simp
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   434
apply (subst strictify_cancel)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   435
apply (simp add: cfun_map_def deflation_strict 1 2)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   436
apply (simp add: cfun_map_def deflation.idem 1 2)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   437
apply (simp add: sfun_below_iff)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   438
apply (subst strictify_cancel)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   439
apply (simp add: cfun_map_def deflation_strict 1 2)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   440
apply (rule deflation.below)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   441
apply (rule deflation_cfun_map [OF 1 2])
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   442
done
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   443
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   444
lemma finite_deflation_sfun_map:
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   445
  assumes 1: "finite_deflation d1"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   446
  assumes 2: "finite_deflation d2"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   447
  shows "finite_deflation (sfun_map\<cdot>d1\<cdot>d2)"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   448
proof (intro finite_deflation_intro)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   449
  interpret d1: finite_deflation d1 by fact
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   450
  interpret d2: finite_deflation d2 by fact
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   451
  have "deflation d1" and "deflation d2" by fact+
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   452
  thus "deflation (sfun_map\<cdot>d1\<cdot>d2)" by (rule deflation_sfun_map)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   453
  from 1 2 have "finite_deflation (cfun_map\<cdot>d1\<cdot>d2)"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   454
    by (rule finite_deflation_cfun_map)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   455
  then have "finite {f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f = f}"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   456
    by (rule finite_deflation.finite_fixes)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   457
  moreover have "inj (\<lambda>f. sfun_rep\<cdot>f)"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   458
    by (rule inj_onI, simp add: sfun_eq_iff)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   459
  ultimately have "finite ((\<lambda>f. sfun_rep\<cdot>f) -` {f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f = f})"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   460
    by (rule finite_vimageI)
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   461
  then show "finite {f. sfun_map\<cdot>d1\<cdot>d2\<cdot>f = f}"
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   462
    unfolding sfun_map_def sfun_eq_iff
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   463
    by (simp add: strictify_cancel
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61169
diff changeset
   464
         deflation_strict \<open>deflation d1\<close> \<open>deflation d2\<close>)
40592
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   465
qed
f432973ce0f6 move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents: 40502
diff changeset
   466
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
diff changeset
   467
end