| 67224 |      1 | (*  Title:    HOL/ex/Conditional_Parametricity_Examples.thy
 | 
|  |      2 |     Author:   Jan Gilcher, Andreas Lochbihler, Dmitriy Traytel, ETH Zürich
 | 
|  |      3 | 
 | 
|  |      4 | Examples for the parametric_constant command
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
|  |      7 | theory Conditional_Parametricity_Examples
 | 
|  |      8 |   imports "HOL-Library.Conditional_Parametricity"
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | definition "bar x xs = rev (x # xs)"
 | 
|  |     12 | parametric_constant bar_def
 | 
|  |     13 | 
 | 
|  |     14 | definition bar2 where "bar2 = bar"
 | 
|  |     15 | parametric_constant bar2_def
 | 
|  |     16 | 
 | 
|  |     17 | parametric_constant bar_thm[transfer_rule]: bar_def
 | 
|  |     18 | parametric_constant bar2_thm1: bar2_def
 | 
|  |     19 | 
 | 
|  |     20 | definition "t1 y x = zip x y"
 | 
|  |     21 | parametric_constant t1_thm: t1_def
 | 
|  |     22 | 
 | 
|  |     23 | definition "t2 f x = f (rev x)"
 | 
|  |     24 | parametric_constant t2_thm: t2_def
 | 
|  |     25 | 
 | 
|  |     26 | definition "t3 xs = rev (rev (xs :: 'b list))"
 | 
|  |     27 | parametric_constant t3_thm: t3_def
 | 
|  |     28 | 
 | 
|  |     29 | definition "t4 f x = rev (f x (f x (rev x)))"
 | 
|  |     30 | parametric_constant t4_thm: t4_def
 | 
|  |     31 | 
 | 
|  |     32 | definition "t5 x y = zip x (rev y)"
 | 
|  |     33 | parametric_constant t5_thm: t5_def
 | 
|  |     34 | 
 | 
|  |     35 | (* Conditional Parametricity*)
 | 
|  |     36 | 
 | 
|  |     37 | definition "t6_1 x y = inf y x"
 | 
|  |     38 | parametric_constant t6_1_thm: t6_1_def
 | 
|  |     39 | 
 | 
|  |     40 | definition "t6_2 x y = sup y x"
 | 
|  |     41 | parametric_constant t6_2_thm: t6_2_def
 | 
|  |     42 | 
 | 
|  |     43 | definition "t6_3 x z y = sup (inf x y) z"
 | 
|  |     44 | parametric_constant t6_3_thm: t6_3_def
 | 
|  |     45 | 
 | 
|  |     46 | definition "t6_4 x xs y = map (sup (inf y x)) xs"
 | 
|  |     47 | parametric_constant t6_4_thm: t6_4_def
 | 
|  |     48 | 
 | 
|  |     49 | definition "t7 x y = (y = x)"
 | 
|  |     50 | parametric_constant t7_thm: t7_def
 | 
|  |     51 | 
 | 
|  |     52 | definition "t8 x y = ((x=y) \<and> (y=x))"
 | 
|  |     53 | parametric_constant t8_thm: t8_def
 | 
|  |     54 | 
 | 
|  |     55 | (* Definition via primrec*)
 | 
|  |     56 | primrec delete where
 | 
|  |     57 |   "delete _ [] = []"
 | 
|  |     58 | | "delete x (y # ys) = (if x = y then ys else y # (delete x ys))"
 | 
|  |     59 | parametric_constant delete_thm: delete_def
 | 
|  |     60 | 
 | 
|  |     61 | definition "foo f x y = (if f x = f y then x else sup y x)"
 | 
|  |     62 | parametric_constant foo_parametricity: foo_def
 | 
|  |     63 | 
 | 
| 67399 |     64 | end
 |