Quot/Examples/SigmaEx.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 20 Jan 2010 16:50:31 +0100
changeset 905 51e5cc3793d2
child 908 1bf4337919d3
permissions -rw-r--r--
Added the Sigma Calculus example
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
905
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     1
theory SigmaEx
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     2
imports Nominal "../QuotMain" "../QuotList" "../QuotProd"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     3
begin
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     4
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     5
atom_decl name
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     6
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     7
datatype robj =
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     8
  rVar "name"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
     9
| rObj "(string \<times> rmethod) list"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    10
| rInv "robj" "string"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    11
| rUpd "robj" "string" "rmethod"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    12
and rmethod =
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    13
  rSig "name" "robj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    14
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    15
(* Need to fix it, just 2 random rules *)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    16
inductive
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    17
    alpha_obj :: "robj \<Rightarrow> robj \<Rightarrow> bool" ("_ \<approx>o _" [100, 100] 100)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    18
and alpha_method :: "rmethod \<Rightarrow> rmethod \<Rightarrow> bool" ("_ \<approx>m _" [100, 100] 100)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    19
where
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    20
  a1: "a = b \<Longrightarrow> (rVar a) \<approx>o (rVar b)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    21
| a2: "a = b \<and> c = d \<Longrightarrow> rSig a c \<approx>m rSig b d"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    22
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    23
lemma alpha_equivps:
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    24
  shows "equivp alpha_obj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    25
  and   "equivp alpha_method"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    26
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    27
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    28
quotient_type
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    29
    obj = robj / alpha_obj
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    30
and method = rmethod / alpha_method
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    31
  by (auto intro: alpha_equivps)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    32
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    33
quotient_definition
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    34
  "Var :: name \<Rightarrow> obj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    35
as
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    36
  "rVar"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    37
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    38
quotient_definition
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    39
  "Obj :: (string \<times> method) list \<Rightarrow> obj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    40
as
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    41
  "rObj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    42
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    43
quotient_definition
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    44
  "Inv :: obj \<Rightarrow> string \<Rightarrow> obj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    45
as
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    46
  "rInv"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    47
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    48
quotient_definition
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    49
  "Upd :: obj \<Rightarrow> string \<Rightarrow> method \<Rightarrow> obj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    50
as
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    51
  "rUpd"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    52
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    53
quotient_definition
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    54
  "Sig :: name \<Rightarrow> obj \<Rightarrow> method"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    55
as
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    56
  "rSig"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    57
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    58
overloading
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    59
  perm_obj \<equiv> "perm :: 'x prm \<Rightarrow> obj \<Rightarrow> obj" (unchecked)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    60
  perm_method   \<equiv> "perm :: 'x prm \<Rightarrow> method \<Rightarrow> method" (unchecked)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    61
begin
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    62
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    63
quotient_definition
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    64
  "perm_obj :: 'x prm \<Rightarrow> obj \<Rightarrow> obj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    65
as
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    66
  "(perm::'x prm \<Rightarrow> robj \<Rightarrow> robj)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    67
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    68
quotient_definition
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    69
  "perm_method :: 'x prm \<Rightarrow> method \<Rightarrow> method"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    70
as
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    71
  "(perm::'x prm \<Rightarrow> rmethod \<Rightarrow> rmethod)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    72
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    73
end
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    74
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    75
lemma tolift:
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    76
"\<forall> fvar.
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    77
 \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    78
 \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    79
 \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    80
 \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    81
 \<forall> fnil.
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    82
 \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    83
 \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    84
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    85
 \<exists> (hom_o\<Colon>robj \<Rightarrow> 'a, hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b, hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c, hom_m\<Colon>rmethod \<Rightarrow> 'd)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    86
 \<in> Respects
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    87
   (prod_rel (alpha_obj ===> op =)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    88
     (prod_rel (list_rel (prod_rel (op =) alpha_method) ===> op =)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    89
       (prod_rel ((prod_rel (op =) alpha_method) ===> op =)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    90
         (alpha_method ===> op =)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    91
       )
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    92
     )
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    93
   ).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    94
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    95
(
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    96
 (\<forall>x. hom_o (rVar x) = fvar x) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    97
 (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    98
 (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
    99
 (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   100
 (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   101
 (hom_d [] = fnil) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   102
 (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   103
 (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   104
)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   105
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   106
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   107
syntax
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   108
  "_expttrn"        :: "pttrn => bool => bool"      ("(3\<exists>\<exists> _./ _)" [0, 10] 10)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   109
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   110
translations
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   111
  "\<exists>\<exists> x. P"   == "Ex (%x. P)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   112
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   113
lemma split_rsp[quot_respect]: "((R1 ===> R2 ===> op =) ===> (prod_rel R1 R2) ===> op =) split split"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   114
by auto
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   115
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   116
lemma rvar_rsp[quot_respect]: "(op = ===> alpha_obj) rVar rVar"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   117
by (simp add: a1)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   118
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   119
lemma robj_rsp[quot_respect]: "(list_rel (prod_rel op = alpha_method) ===> alpha_obj) rObj rObj"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   120
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   121
lemma rinv_rsp[quot_respect]: "(alpha_obj ===> op = ===> alpha_obj) rInv rInv"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   122
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   123
lemma rupd_rsp[quot_respect]: "(alpha_obj ===> op = ===> alpha_method ===> alpha_obj) rUpd rUpd"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   124
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   125
lemma rsig_rsp[quot_respect]: "(op = ===> alpha_obj ===> alpha_method) rSig rSig"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   126
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   127
lemma operm_rsp[quot_respect]: "(op = ===> alpha_obj ===> alpha_obj) op \<bullet> op \<bullet>"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   128
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   129
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   130
lemma liftd: "
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   131
\<exists>\<exists>(hom_o, (hom_d, (hom_e, hom_m))).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   132
(
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   133
 (\<forall>x. hom_o (Var x) = fvar x) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   134
 (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   135
 (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   136
 (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   137
 (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   138
 (hom_d [] = fnil) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   139
 (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   140
 (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   141
)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   142
apply (lifting tolift)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   143
apply (regularize)
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   144
prefer 2
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   145
apply cleaning
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   146
apply simp
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   147
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   148
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   149
lemma tolift':
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   150
"\<forall> fvar.
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   151
 \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   152
 \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   153
 \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   154
 \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
 \<forall> fnil.
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   156
 \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   157
 \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   158
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   159
 \<exists> hom_o\<Colon>robj \<Rightarrow> 'a \<in> Respects (alpha_obj ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   160
 \<exists> hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b \<in> Respects (list_rel (prod_rel (op =) alpha_method) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   161
 \<exists> hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c \<in> Respects ((prod_rel (op =) alpha_method) ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   162
 \<exists> hom_m\<Colon>rmethod \<Rightarrow> 'd \<in> Respects (alpha_method ===> op =).
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   163
(
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   164
 (\<forall>x. hom_o (rVar x) = fvar x) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   165
 (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   166
 (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   167
 (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   168
 (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   169
 (hom_d [] = fnil) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   170
 (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   171
 (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   172
)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   173
sorry
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   174
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   175
lemma liftd': "
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   176
\<exists>hom_o. \<exists>hom_d. \<exists>hom_e. \<exists>hom_m.
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   177
(
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   178
 (\<forall>x. hom_o (Var x) = fvar x) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   179
 (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   180
 (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   181
 (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   182
 (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   183
 (hom_d [] = fnil) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   184
 (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   185
 (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   186
)"
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   187
apply (lifting tolift')
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   188
done
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   189
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   190
end
51e5cc3793d2 Added the Sigma Calculus example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   191