Anforderungen  |   Konzepte  |   Entwurf  |   Entwicklung  |   Qualitätssicherung  |   Lebenszyklus  |   Steuerung
 
 
 
 


Quellcode-Bibliothek

© Kompilation durch diese Firma

[Weder Korrektheit noch Funktionsfähigkeit der Software werden zugesichert.]

Datei: Map_Functions.thy   Sprache: Isabelle

Original von: Isabelle©

(*  Title:      HOL/HOLCF/Map_Functions.thy
    Author:     Brian Huffman
*)


section \<open>Map functions for various types\<close>

theory Map_Functions
  imports Deflation Sprod Ssum Sfun Up
begin

subsection \<open>Map operator for continuous function space\<close>

default_sort cpo

definition cfun_map :: "('b \ 'a) \ ('c \ 'd) \ ('a \ 'c) \ ('b \ 'd)"
  where "cfun_map = (\ a b f x. b\(f\(a\x)))"

lemma cfun_map_beta [simp]: "cfun_map\a\b\f\x = b\(f\(a\x))"
  by (simp add: cfun_map_def)

lemma cfun_map_ID: "cfun_map\ID\ID = ID"
  by (simp add: cfun_eq_iff)

lemma cfun_map_map: "cfun_map\f1\g1\(cfun_map\f2\g2\p) = cfun_map\(\ x. f2\(f1\x))\(\ x. g1\(g2\x))\p"
  by (rule cfun_eqI) simp

lemma ep_pair_cfun_map:
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
  shows "ep_pair (cfun_map\p1\e2) (cfun_map\e1\p2)"
proof
  interpret e1p1: ep_pair e1 p1 by fact
  interpret e2p2: ep_pair e2 p2 by fact
  show "cfun_map\e1\p2\(cfun_map\p1\e2\f) = f" for f
    by (simp add: cfun_eq_iff)
  show "cfun_map\p1\e2\(cfun_map\e1\p2\g) \ g" for g
    apply (rule cfun_belowI, simp)
    apply (rule below_trans [OF e2p2.e_p_below])
    apply (rule monofun_cfun_arg)
    apply (rule e1p1.e_p_below)
    done
qed

lemma deflation_cfun_map:
  assumes "deflation d1" and "deflation d2"
  shows "deflation (cfun_map\d1\d2)"
proof
  interpret d1: deflation d1 by fact
  interpret d2: deflation d2 by fact
  fix f
  show "cfun_map\d1\d2\(cfun_map\d1\d2\f) = cfun_map\d1\d2\f"
    by (simp add: cfun_eq_iff d1.idem d2.idem)
  show "cfun_map\d1\d2\f \ f"
    apply (rule cfun_belowI, simp)
    apply (rule below_trans [OF d2.below])
    apply (rule monofun_cfun_arg)
    apply (rule d1.below)
    done
qed

lemma finite_range_cfun_map:
  assumes a: "finite (range (\x. a\x))"
  assumes b: "finite (range (\y. b\y))"
  shows "finite (range (\f. cfun_map\a\b\f))" (is "finite (range ?h)")
proof (rule finite_imageD)
  let ?f = "\g. range (\x. (a\x, g\x))"
  show "finite (?f ` range ?h)"
  proof (rule finite_subset)
    let ?B = "Pow (range (\x. a\x) \ range (\y. b\y))"
    show "?f ` range ?h \ ?B"
      by clarsimp
    show "finite ?B"
      by (simp add: a b)
  qed
  show "inj_on ?f (range ?h)"
  proof (rule inj_onI, rule cfun_eqI, clarsimp)
    fix x f g
    assume "range (\x. (a\x, b\(f\(a\x)))) = range (\x. (a\x, b\(g\(a\x))))"
    then have "range (\x. (a\x, b\(f\(a\x)))) \ range (\x. (a\x, b\(g\(a\x))))"
      by (rule equalityD1)
    then have "(a\x, b\(f\(a\x))) \ range (\x. (a\x, b\(g\(a\x))))"
      by (simp add: subset_eq)
    then obtain y where "(a\x, b\(f\(a\x))) = (a\y, b\(g\(a\y)))"
      by (rule rangeE)
    then show "b\(f\(a\x)) = b\(g\(a\x))"
      by clarsimp
  qed
qed

lemma finite_deflation_cfun_map:
  assumes "finite_deflation d1" and "finite_deflation d2"
  shows "finite_deflation (cfun_map\d1\d2)"
proof (rule finite_deflation_intro)
  interpret d1: finite_deflation d1 by fact
  interpret d2: finite_deflation d2 by fact
  from d1.deflation_axioms d2.deflation_axioms show "deflation (cfun_map\d1\d2)"
    by (rule deflation_cfun_map)
  have "finite (range (\f. cfun_map\d1\d2\f))"
    using d1.finite_range d2.finite_range
    by (rule finite_range_cfun_map)
  then show "finite {f. cfun_map\d1\d2\f = f}"
    by (rule finite_range_imp_finite_fixes)
qed

text \<open>Finite deflations are compact elements of the function space\<close>

lemma finite_deflation_imp_compact: "finite_deflation d \ compact d"
  apply (frule finite_deflation_imp_deflation)
  apply (subgoal_tac "compact (cfun_map\d\d\d)")
   apply (simp add: cfun_map_def deflation.idem eta_cfun)
  apply (rule finite_deflation.compact)
  apply (simp only: finite_deflation_cfun_map)
  done


subsection \<open>Map operator for product type\<close>

definition prod_map :: "('a \ 'b) \ ('c \ 'd) \ 'a \ 'c \ 'b \ 'd"
  where "prod_map = (\ f g p. (f\(fst p), g\(snd p)))"

lemma prod_map_Pair [simp]: "prod_map\f\g\(x, y) = (f\x, g\y)"
  by (simp add: prod_map_def)

lemma prod_map_ID: "prod_map\ID\ID = ID"
  by (auto simp: cfun_eq_iff)

lemma prod_map_map: "prod_map\f1\g1\(prod_map\f2\g2\p) = prod_map\(\ x. f1\(f2\x))\(\ x. g1\(g2\x))\p"
  by (induct p) simp

lemma ep_pair_prod_map:
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
  shows "ep_pair (prod_map\e1\e2) (prod_map\p1\p2)"
proof
  interpret e1p1: ep_pair e1 p1 by fact
  interpret e2p2: ep_pair e2 p2 by fact
  show "prod_map\p1\p2\(prod_map\e1\e2\x) = x" for x
    by (induct x) simp
  show "prod_map\e1\e2\(prod_map\p1\p2\y) \ y" for y
    by (induct y) (simp add: e1p1.e_p_below e2p2.e_p_below)
qed

lemma deflation_prod_map:
  assumes "deflation d1" and "deflation d2"
  shows "deflation (prod_map\d1\d2)"
proof
  interpret d1: deflation d1 by fact
  interpret d2: deflation d2 by fact
  fix x
  show "prod_map\d1\d2\(prod_map\d1\d2\x) = prod_map\d1\d2\x"
    by (induct x) (simp add: d1.idem d2.idem)
  show "prod_map\d1\d2\x \ x"
    by (induct x) (simp add: d1.below d2.below)
qed

lemma finite_deflation_prod_map:
  assumes "finite_deflation d1" and "finite_deflation d2"
  shows "finite_deflation (prod_map\d1\d2)"
proof (rule finite_deflation_intro)
  interpret d1: finite_deflation d1 by fact
  interpret d2: finite_deflation d2 by fact
  from d1.deflation_axioms d2.deflation_axioms show "deflation (prod_map\d1\d2)"
    by (rule deflation_prod_map)
  have "{p. prod_map\d1\d2\p = p} \ {x. d1\x = x} \ {y. d2\y = y}"
    by auto
  then show "finite {p. prod_map\d1\d2\p = p}"
    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
qed


subsection \<open>Map function for lifted cpo\<close>

definition u_map :: "('a \ 'b) \ 'a u \ 'b u"
  where "u_map = (\ f. fup\(up oo f))"

lemma u_map_strict [simp]: "u_map\f\\ = \"
  by (simp add: u_map_def)

lemma u_map_up [simp]: "u_map\f\(up\x) = up\(f\x)"
  by (simp add: u_map_def)

lemma u_map_ID: "u_map\ID = ID"
  by (simp add: u_map_def cfun_eq_iff eta_cfun)

lemma u_map_map: "u_map\f\(u_map\g\p) = u_map\(\ x. f\(g\x))\p"
  by (induct p) simp_all

lemma u_map_oo: "u_map\(f oo g) = u_map\f oo u_map\g"
  by (simp add: cfcomp1 u_map_map eta_cfun)

lemma ep_pair_u_map: "ep_pair e p \ ep_pair (u_map\e) (u_map\p)"
  apply standard
  subgoal for x by (cases x) (simp_all add: ep_pair.e_inverse)
  subgoal for y by (cases y) (simp_all add: ep_pair.e_p_below)
  done

lemma deflation_u_map: "deflation d \ deflation (u_map\d)"
  apply standard
  subgoal for x by (cases x) (simp_all add: deflation.idem)
  subgoal for x by (cases x) (simp_all add: deflation.below)
  done

lemma finite_deflation_u_map:
  assumes "finite_deflation d"
  shows "finite_deflation (u_map\d)"
proof (rule finite_deflation_intro)
  interpret d: finite_deflation d by fact
  from d.deflation_axioms show "deflation (u_map\d)"
    by (rule deflation_u_map)
  have "{x. u_map\d\x = x} \ insert \ ((\x. up\x) ` {x. d\x = x})"
    by (rule subsetI, case_tac x, simp_all)
  then show "finite {x. u_map\d\x = x}"
    by (rule finite_subset) (simp add: d.finite_fixes)
qed


subsection \<open>Map function for strict products\<close>

default_sort pcpo

definition sprod_map :: "('a \ 'b) \ ('c \ 'd) \ 'a \ 'c \ 'b \ 'd"
  where "sprod_map = (\ f g. ssplit\(\ x y. (:f\x, g\y:)))"

lemma sprod_map_strict [simp]: "sprod_map\a\b\\ = \"
  by (simp add: sprod_map_def)

lemma sprod_map_spair [simp]: "x \ \ \ y \ \ \ sprod_map\f\g\(:x, y:) = (:f\x, g\y:)"
  by (simp add: sprod_map_def)

lemma sprod_map_spair': "f\\ = \ \ g\\ = \ \ sprod_map\f\g\(:x, y:) = (:f\x, g\y:)"
  by (cases "x = \ \ y = \") auto

lemma sprod_map_ID: "sprod_map\ID\ID = ID"
  by (simp add: sprod_map_def cfun_eq_iff eta_cfun)

lemma sprod_map_map:
  "\f1\\ = \; g1\\ = \\ \
    sprod_map\<cdot>f1\<cdot>g1\<cdot>(sprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
     sprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
proof (induct p)
  case bottom
  then show ?case by simp
next
  case (spair x y)
  then show ?case
    apply (cases "f2\x = \", simp)
    apply (cases "g2\y = \", simp)
    apply simp
    done
qed

lemma ep_pair_sprod_map:
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
  shows "ep_pair (sprod_map\e1\e2) (sprod_map\p1\p2)"
proof
  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
  show "sprod_map\p1\p2\(sprod_map\e1\e2\x) = x" for x
    by (induct x) simp_all
  show "sprod_map\e1\e2\(sprod_map\p1\p2\y) \ y" for y
  proof (induct y)
    case bottom
    then show ?case by simp
  next
    case (spair x y)
    then show ?case
      apply simp
      apply (cases "p1\x = \", simp, cases "p2\y = \", simp)
      apply (simp add: monofun_cfun e1p1.e_p_below e2p2.e_p_below)
      done
  qed
qed

lemma deflation_sprod_map:
  assumes "deflation d1" and "deflation d2"
  shows "deflation (sprod_map\d1\d2)"
proof
  interpret d1: deflation d1 by fact
  interpret d2: deflation d2 by fact
  fix x
  show "sprod_map\d1\d2\(sprod_map\d1\d2\x) = sprod_map\d1\d2\x"
  proof (induct x)
    case bottom
    then show ?case by simp
  next
    case (spair x y)
    then show ?case
      apply (cases "d1\x = \", simp, cases "d2\y = \", simp)
      apply (simp add: d1.idem d2.idem)
      done
  qed
  show "sprod_map\d1\d2\x \ x"
  proof (induct x)
    case bottom
    then show ?case by simp
  next
    case spair
    then show ?case by (simp add: monofun_cfun d1.below d2.below)
  qed
qed

lemma finite_deflation_sprod_map:
  assumes "finite_deflation d1" and "finite_deflation d2"
  shows "finite_deflation (sprod_map\d1\d2)"
proof (rule finite_deflation_intro)
  interpret d1: finite_deflation d1 by fact
  interpret d2: finite_deflation d2 by fact
  from d1.deflation_axioms d2.deflation_axioms show "deflation (sprod_map\d1\d2)"
    by (rule deflation_sprod_map)
  have "{x. sprod_map\d1\d2\x = x} \
      insert \<bottom> ((\<lambda>(x, y). (:x, y:)) ` ({x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}))"
    by (rule subsetI, case_tac x, auto simp add: spair_eq_iff)
  then show "finite {x. sprod_map\d1\d2\x = x}"
    by (rule finite_subset) (simp add: d1.finite_fixes d2.finite_fixes)
qed


subsection \<open>Map function for strict sums\<close>

definition ssum_map :: "('a \ 'b) \ ('c \ 'd) \ 'a \ 'c \ 'b \ 'd"
  where "ssum_map = (\ f g. sscase\(sinl oo f)\(sinr oo g))"

lemma ssum_map_strict [simp]: "ssum_map\f\g\\ = \"
  by (simp add: ssum_map_def)

lemma ssum_map_sinl [simp]: "x \ \ \ ssum_map\f\g\(sinl\x) = sinl\(f\x)"
  by (simp add: ssum_map_def)

lemma ssum_map_sinr [simp]: "x \ \ \ ssum_map\f\g\(sinr\x) = sinr\(g\x)"
  by (simp add: ssum_map_def)

lemma ssum_map_sinl': "f\\ = \ \ ssum_map\f\g\(sinl\x) = sinl\(f\x)"
  by (cases "x = \") simp_all

lemma ssum_map_sinr': "g\\ = \ \ ssum_map\f\g\(sinr\x) = sinr\(g\x)"
  by (cases "x = \") simp_all

lemma ssum_map_ID: "ssum_map\ID\ID = ID"
  by (simp add: ssum_map_def cfun_eq_iff eta_cfun)

lemma ssum_map_map:
  "\f1\\ = \; g1\\ = \\ \
    ssum_map\<cdot>f1\<cdot>g1\<cdot>(ssum_map\<cdot>f2\<cdot>g2\<cdot>p) =
     ssum_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
proof (induct p)
  case bottom
  then show ?case by simp
next
  case (sinl x)
  then show ?case by (cases "f2\x = \") simp_all
next
  case (sinr y)
  then show ?case by (cases "g2\y = \") simp_all
qed

lemma ep_pair_ssum_map:
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
  shows "ep_pair (ssum_map\e1\e2) (ssum_map\p1\p2)"
proof
  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
  show "ssum_map\p1\p2\(ssum_map\e1\e2\x) = x" for x
    by (induct x) simp_all
  show "ssum_map\e1\e2\(ssum_map\p1\p2\y) \ y" for y
  proof (induct y)
    case bottom
    then show ?case by simp
  next
    case (sinl x)
    then show ?case by (cases "p1\x = \") (simp_all add: e1p1.e_p_below)
  next
    case (sinr y)
    then show ?case by (cases "p2\y = \") (simp_all add: e2p2.e_p_below)
  qed
qed

lemma deflation_ssum_map:
  assumes "deflation d1" and "deflation d2"
  shows "deflation (ssum_map\d1\d2)"
proof
  interpret d1: deflation d1 by fact
  interpret d2: deflation d2 by fact
  fix x
  show "ssum_map\d1\d2\(ssum_map\d1\d2\x) = ssum_map\d1\d2\x"
  proof (induct x)
    case bottom
    then show ?case by simp
  next
    case (sinl x)
    then show ?case by (cases "d1\x = \") (simp_all add: d1.idem)
  next
    case (sinr y)
    then show ?case by (cases "d2\y = \") (simp_all add: d2.idem)
  qed
  show "ssum_map\d1\d2\x \ x"
  proof (induct x)
    case bottom
    then show ?case by simp
  next
    case (sinl x)
    then show ?case by (cases "d1\x = \") (simp_all add: d1.below)
  next
    case (sinr y)
    then show ?case by (cases "d2\y = \") (simp_all add: d2.below)
  qed
qed

lemma finite_deflation_ssum_map:
  assumes "finite_deflation d1" and "finite_deflation d2"
  shows "finite_deflation (ssum_map\d1\d2)"
proof (rule finite_deflation_intro)
  interpret d1: finite_deflation d1 by fact
  interpret d2: finite_deflation d2 by fact
  from d1.deflation_axioms d2.deflation_axioms show "deflation (ssum_map\d1\d2)"
    by (rule deflation_ssum_map)
  have "{x. ssum_map\d1\d2\x = x} \
        (\<lambda>x. sinl\<cdot>x) ` {x. d1\<cdot>x = x} \<union>
        (\<lambda>x. sinr\<cdot>x) ` {x. d2\<cdot>x = x} \<union> {\<bottom>}"
    by (rule subsetI, case_tac x, simp_all)
  then show "finite {x. ssum_map\d1\d2\x = x}"
    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
qed


subsection \<open>Map operator for strict function space\<close>

definition sfun_map :: "('b \ 'a) \ ('c \ 'd) \ ('a \! 'c) \ ('b \! 'd)"
  where "sfun_map = (\ a b. sfun_abs oo cfun_map\a\b oo sfun_rep)"

lemma sfun_map_ID: "sfun_map\ID\ID = ID"
  by (simp add: sfun_map_def cfun_map_ID cfun_eq_iff)

lemma sfun_map_map:
  assumes "f2\\ = \" and "g2\\ = \"
  shows "sfun_map\f1\g1\(sfun_map\f2\g2\p) =
    sfun_map\<cdot>(\<Lambda> x. f2\<cdot>(f1\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
  by (simp add: sfun_map_def cfun_eq_iff strictify_cancel assms cfun_map_map)

lemma ep_pair_sfun_map:
  assumes 1: "ep_pair e1 p1"
  assumes 2: "ep_pair e2 p2"
  shows "ep_pair (sfun_map\p1\e2) (sfun_map\e1\p2)"
proof
  interpret e1p1: pcpo_ep_pair e1 p1
    unfolding pcpo_ep_pair_def by fact
  interpret e2p2: pcpo_ep_pair e2 p2
    unfolding pcpo_ep_pair_def by fact
  show "sfun_map\e1\p2\(sfun_map\p1\e2\f) = f" for f
    unfolding sfun_map_def
    apply (simp add: sfun_eq_iff strictify_cancel)
    apply (rule ep_pair.e_inverse)
    apply (rule ep_pair_cfun_map [OF 1 2])
    done
  show "sfun_map\p1\e2\(sfun_map\e1\p2\g) \ g" for g
    unfolding sfun_map_def
    apply (simp add: sfun_below_iff strictify_cancel)
    apply (rule ep_pair.e_p_below)
    apply (rule ep_pair_cfun_map [OF 1 2])
    done
qed

lemma deflation_sfun_map:
  assumes 1: "deflation d1"
  assumes 2: "deflation d2"
  shows "deflation (sfun_map\d1\d2)"
  apply (simp add: sfun_map_def)
  apply (rule deflation.intro)
   apply simp
   apply (subst strictify_cancel)
    apply (simp add: cfun_map_def deflation_strict 1 2)
   apply (simp add: cfun_map_def deflation.idem 1 2)
  apply (simp add: sfun_below_iff)
  apply (subst strictify_cancel)
   apply (simp add: cfun_map_def deflation_strict 1 2)
  apply (rule deflation.below)
  apply (rule deflation_cfun_map [OF 1 2])
  done

lemma finite_deflation_sfun_map:
  assumes "finite_deflation d1"
    and "finite_deflation d2"
  shows "finite_deflation (sfun_map\d1\d2)"
proof (intro finite_deflation_intro)
  interpret d1: finite_deflation d1 by fact
  interpret d2: finite_deflation d2 by fact
  from d1.deflation_axioms d2.deflation_axioms show "deflation (sfun_map\d1\d2)"
    by (rule deflation_sfun_map)
  from assms have "finite_deflation (cfun_map\d1\d2)"
    by (rule finite_deflation_cfun_map)
  then have "finite {f. cfun_map\d1\d2\f = f}"
    by (rule finite_deflation.finite_fixes)
  moreover have "inj (\f. sfun_rep\f)"
    by (rule inj_onI) (simp add: sfun_eq_iff)
  ultimately have "finite ((\f. sfun_rep\f) -` {f. cfun_map\d1\d2\f = f})"
    by (rule finite_vimageI)
  with \<open>deflation d1\<close> \<open>deflation d2\<close> show "finite {f. sfun_map\<cdot>d1\<cdot>d2\<cdot>f = f}"
    by (simp add: sfun_map_def sfun_eq_iff strictify_cancel deflation_strict)
qed

end

¤ Dauer der Verarbeitung: 0.24 Sekunden  (vorverarbeitet)  ¤





Download des
Quellennavigators
Download des
sprechenden Kalenders

in der Quellcodebibliothek suchen




Haftungshinweis

Die Informationen auf dieser Webseite wurden nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit, noch Qualität der bereit gestellten Informationen zugesichert.


Bemerkung:

Die farbliche Syntaxdarstellung ist noch experimentell.


Bot Zugriff



                                                                                                                                                                                                                                                                                                                                                                                                     


Neuigkeiten

     Aktuelles
     Motto des Tages

Software

     Produkte
     Quellcodebibliothek

Aktivitäten

     Artikel über Sicherheit
     Anleitung zur Aktivierung von SSL

Muße

     Gedichte
     Musik
     Bilder

Jenseits des Üblichen ....
    

Besucherstatistik

Besucherstatistik