prio/Precedence_ord.thy
author urbanc
Wed, 02 May 2012 13:13:47 +0000
changeset 352 ee58e3d99f8a
parent 320 630754a81bdb
permissions -rw-r--r--
added section about PINTOS and rewritten multi-processor section
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     1
header {* Order on product types *}
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     2
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     3
theory Precedence_ord
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     4
imports Main
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     5
begin
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     6
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     7
datatype precedence = Prc nat nat
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     8
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     9
instantiation precedence :: order
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    10
begin
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    11
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    12
definition
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    13
  precedence_le_def: "x \<le> y \<longleftrightarrow> (case (x, y) of
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    14
                                   (Prc fx sx, Prc fy sy) \<Rightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    15
                                 fx < fy \<or> (fx \<le> fy \<and> sy \<le> sx))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    16
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    17
definition
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    18
  precedence_less_def: "x < y \<longleftrightarrow> (case (x, y) of
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    19
                               (Prc fx sx, Prc fy sy) \<Rightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    20
                                     fx < fy \<or> (fx \<le> fy \<and> sy < sx))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    21
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    22
instance
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    23
proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    24
qed (auto simp: precedence_le_def precedence_less_def 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    25
              intro: order_trans split:precedence.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    26
end
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    27
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    28
instance precedence :: preorder ..
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    29
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    30
instance precedence :: linorder proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    31
qed (auto simp: precedence_le_def precedence_less_def 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    32
              intro: order_trans split:precedence.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    33
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    34
end