Theory Restriction_Spaces-HOLCF

(***********************************************************************************
 * Copyright (c) 2025 Université Paris-Saclay
 *
 * Author: Benoît Ballenghien, Université Paris-Saclay,
           CNRS, ENS Paris-Saclay, LMF
 * Author: Benjamin Puyobro, Université Paris-Saclay,
           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
 * Author: Burkhart Wolff, Université Paris-Saclay,
           CNRS, ENS Paris-Saclay, LMF
 *
 * All rights reserved.
 *
 * Redistribution and use in source and binary forms, with or without
 * modification, are permitted provided that the following conditions are met:
 *
 * * Redistributions of source code must retain the above copyright notice, this
 *
 * * Redistributions in binary form must reproduce the above copyright notice,
 *   this list of conditions and the following disclaimer in the documentation
 *   and/or other materials provided with the distribution.
 *
 * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
 * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
 * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
 * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
 * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
 * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
 * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
 * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
 * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
 * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
 *
 * SPDX-License-Identifier: BSD-2-Clause
 ***********************************************************************************)

(*<*)
theory "Restriction_Spaces-HOLCF"
  imports HOLCF Restriction_Spaces
begin

default_sort type
  (*>*)


section ‹Definitions›

interpretation belowRS : Restriction (↓) (⊑) by unfold_locales
    ― ‹Just to recover constbelowRS.restriction_related_set and
   constbelowRS.restriction_not_related_set.›

class po_restriction_space = restriction + po +
  assumes restriction_0_below [simp] : x  0  y  0
    and mono_restriction_below     : x  y  x  n  y  n
    and ex_not_restriction_below   : ¬ x  y  n. ¬ x  n  y  n

interpretation belowRS : PreorderRestrictionSpace (↓) :: 'a  nat  'a :: po_restriction_space (⊑)
proof unfold_locales
  show x  0  y  0 for x y :: 'a by simp
next
  show x  y  x  n  y  n for x y :: 'a and n
    by (fact mono_restriction_below)
next
  show ¬ x  y  n. ¬ x  n  y  n for x y :: 'a
    by (simp add: ex_not_restriction_below)
next
  show x  y  y  z  x  z for x y z :: 'a by (fact below_trans)
qed


subclass (in po_restriction_space) restriction_space
proof unfold_locales
  show x  0 = y  0 for x y :: 'a by (rule below_antisym) simp_all
next
  show x  y  n. x  n  y  n for x y :: 'a
    by (metis ex_not_restriction_below po_eq_conv)
qed




class cpo_restriction_space = po_restriction_space +
  assumes cpo : chain S  x. range S <<| x


subclass (in cpo_restriction_space) cpo
  by unfold_locales (fact cpo)


class pcpo_restriction_space = cpo_restriction_space +
  assumes least : x. y. x  y

subclass (in pcpo_restriction_space) pcpo
  by unfold_locales (fact least)


interpretation belowRS : Restriction_2_PreorderRestrictionSpace
  (↓) :: 'a :: {restriction, below}  nat  'a (⊑)
  (↓) :: 'b :: po_restriction_space  nat  'b (⊑) ..

text ‹With this we recover constants like constless_eqRS.restriction_shift_on.›

interpretation belowRS : PreorderRestrictionSpace_2_PreorderRestrictionSpace
  (↓) :: 'a :: po_restriction_space  nat  'a (⊑)
  (↓) :: 'b :: po_restriction_space  nat  'b (⊑) ..

text ‹With that we recover theorems like @{thm belowRS.constructive_restriction_restriction}.›

interpretation belowRS : Restriction_2_PreorderRestrictionSpace_2_PreorderRestrictionSpace
  (↓) :: 'a :: {restriction, below}  nat  'a (⊑)
  (↓) :: 'b :: po_restriction_space  nat  'b (⊑)
  (↓) :: 'c :: po_restriction_space  nat  'c (⊑)..

text ‹And with that we recover theorems like @{thm less_eqRS.constructive_on_comp_non_destructive_on}.›



context fixes f :: 'a :: restriction  'b :: po_restriction_space begin
text ‹From @{thm below_antisym} we can obtain stronger lemmas.›

corollary below_restriction_shift_onI :
  (x y n. x  A; y  A; f x  f y; x  n = y  n 
             f x  nat (int n + k)  f y  nat (int n + k))
    restriction_shift_on f k A
  by (simp add: below_antisym restriction_shift_onI)

corollary below_restriction_shiftI :
  (x y n. f x  f y; x  n = y  n 
             f x  nat (int n + k)  f y  nat (int n + k))
    restriction_shift f k
  by (simp add: below_antisym restriction_shiftI)

corollary below_non_too_destructive_onI :
  (x y n. x  A; y  A; f x  f y; x  Suc n = y  Suc n 
             f x  n  f y  n)
    non_too_destructive_on f A
  by (simp add: below_antisym non_too_destructive_onI)

corollary below_non_too_destructiveI :
  (x y n. f x  f y; x  Suc n = y  Suc n  f x  n  f y  n)
    non_too_destructive f
  by (simp add: below_antisym non_too_destructiveI)

corollary below_non_destructive_onI :
  (x y n. n  0; x  A; y  A; f x  f y; x  n = y  n  f x  n  f y  n)
    non_destructive_on f A
  by (simp add: below_antisym non_destructive_onI)

corollary below_non_destructiveI :
  (x y n. n  0; f x  f y; x  n = y  n  f x  n  f y  n)
    non_destructive f
  by (simp add: below_antisym non_destructiveI)

corollary below_constructive_onI :
  (x y n. x  A; y  A; f x  f y; x  n = y  n  f x  Suc n  f y  Suc n)
    constructive_on f A
  by (simp add: below_antisym constructive_onI)

corollary below_constructiveI :
  (x y n. f x  f y; x  n = y  n  f x  Suc n  f y  Suc n)
    constructive f
  by (simp add: below_antisym constructiveI)

end



section ‹Equality of Fixed-Point Operators›

lemma restriction_fix_is_fix :
  (υ X. f X) = (μ X. f X) if cont f constructive f
for f :: 'a :: {pcpo_restriction_space, complete_restriction_space}  'a
proof (rule restriction_fix_unique)
  show constructive f by (fact constructive f)
next
  show f (μ x. f x) = (μ x. f x) by (metis def_cont_fix_eq cont f)
qed



section ‹Product›

instance prod :: (po_restriction_space, po_restriction_space) po_restriction_space
proof intro_classes
  show p  0  q  0 for p q :: 'a × 'b
    by (metis below_refl restriction_0_related)
next
  show p  q  p  n  q  n for p q :: 'a × 'b and n
    by (simp add: below_prod_def restriction_prod_def mono_restriction_below)
next
  show ¬ p  q  n. ¬ p  n  q  n for p q :: 'a × 'b
    by (simp add: below_prod_def restriction_prod_def)
      (metis ex_not_restriction_below)
qed


instance prod :: (cpo_restriction_space, cpo_restriction_space) cpo_restriction_space
  using is_lub_prod by intro_classes blast

instance prod :: (pcpo_restriction_space, pcpo_restriction_space) pcpo_restriction_space
  by (intro_classes) (simp add: pcpo_class.least)


section ‹Functions›

instance fun :: (type, po_restriction_space) po_restriction_space
proof intro_classes
  show f  0  g  0 for f g :: 'a  'b
    by (metis below_refl restriction_0_related)
next
  show f  g  f  n  g  n for f g :: 'a  'b and n
    by (simp add: fun_below_iff mono_restriction_below restriction_fun_def)
next
  show ¬ f  g  n. ¬ f  n  g  n for f g :: 'a  'b
    by (metis belowRS.all_ge_restriction_related_iff_related fun_below_iff restriction_fun_def)
qed

instance fun :: (type, cpo_restriction_space) cpo_restriction_space
  by intro_classes (simp add: cpo_class.cpo)

instance fun :: (type, pcpo_restriction_space) pcpo_restriction_space
  by intro_classes (simp add: pcpo_class.least)


(*<*)
end
  (*>*)