├── Makefile ├── Readme.md ├── _CoqProject ├── src ├── g_typing_flags.ml4 ├── typing_flags_main.ml ├── typing_flags_main.mli └── typing_flags_plugin.mlpack └── theories ├── BreadthFirst.v ├── Demo.v ├── Loader.v └── SameFringe.v /Makefile: -------------------------------------------------------------------------------- 1 | all: Makefile.coq 2 | $(MAKE) -f Makefile.coq 3 | 4 | .PHONY: all install html clean mrproper 5 | 6 | install: Makefile.coq 7 | $(MAKE) -f Makefile.coq install 8 | 9 | html: all 10 | $(MAKE) -f Makefile.coq html 11 | 12 | clean: Makefile.coq 13 | $(MAKE) -f Makefile.coq clean 14 | 15 | mrproper: clean 16 | rm -f theories/.*.aux 17 | rm -f Makefile.coq.conf 18 | rm -f Makefile.coq 19 | 20 | Makefile.coq: _CoqProject 21 | coq_makefile -f _CoqProject -o Makefile.coq 22 | 23 | .merlin: Makefile.coq 24 | $(MAKE) -f Makefile.coq .merlin 25 | -------------------------------------------------------------------------------- /Readme.md: -------------------------------------------------------------------------------- 1 | _Typing Flags_ is a Coq plugin to disable positivity check, guard check and termination check. 2 | 3 | Disabling those checks leads of course to inconsistencies, use it at your own risks ... 4 | 5 | This plugin is not perfect but it does the job. A better version will be integrated in Coq 8.11. 6 | 7 | Tested with coq 8.6, 8.7, 8.8, 8.9. 8 | 9 | 10 | 11 | Usage 12 | ===== 13 | 14 | Load the plugin with `From TypingFlags Require Import Loader.`. 15 | 16 | Then, the plugin provides the following commands: 17 | 18 | - `Set Type In Type` 19 | Globally enable "type in type", meaning that all "Type@{i}" are convertible. Prop and Set remain distinct. 20 | 21 | - `Unset Type In Type` 22 | Disable "type in type". 23 | 24 | - `Unset Guard Checking` 25 | Disable the positivity checking of inductive types and the termination checking of fixpoints. 26 | 27 | - `Set Guard Checking` 28 | (Re)enable positivity and termination checking. Note that a fixpoint defined unguarded may be re-type checked if it is unfolded for a reason or another (Does anybody have a nice example?) 29 | 30 | - `Print Typing Flags` 31 | Print the status of typing flags. 32 | 33 | The `Print Assumptions` command keeps track of which definition and inductives are unsafe. The command `About` also prints safety datas. 34 | 35 | 36 | Installation 37 | ============ 38 | 39 | With OPAM 40 | --------- 41 | 42 | `opam install coq-typing-flags` 43 | 44 | Magic, isn't it? Thanks to @clarus for the package. 45 | 46 | Manually 47 | -------- 48 | 49 | Run `make` to compile the plugin, it needs to have `coqc` available. 50 | 51 | Then, provided the repository theories of the plugin is in your load path, load it with: 52 | 53 | `Require Import TypingFlags.Loader.` 54 | 55 | To copy the plugin in the contrib directory of your local installation of Coq, and thus make it always available simply run: 56 | 57 | `make install` 58 | 59 | 60 | Examples 61 | ======== 62 | 63 | - `theories/Demo.v` 64 | 65 | - `theories/BreadthFirst.v` 66 | Implementation of Martin Hofmann breadthfirst algorithm, presented by Ralph Matthes in TYPES 2018. 67 | 68 | 69 | Acknowledgments 70 | =============== 71 | 72 | This plugin does not implement anything and only make accessible to the user some hidden functionalities of Coq. The implementation is due to Coq developers and in particular Arnaud Spiwack. 73 | 74 | References: https://github.com/coq/coq/pull/79, https://github.com/coq/coq/pull/7651 75 | 76 | 77 | Feedback and other use cases welcome. 78 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R theories/ TypingFlags 2 | -I src 3 | 4 | theories/Loader.v 5 | theories/Demo.v 6 | theories/BreadthFirst.v 7 | theories/SameFringe.v 8 | 9 | src/typing_flags_main.ml 10 | src/typing_flags_main.mli 11 | src/g_typing_flags.ml4 12 | src/typing_flags_plugin.mlpack 13 | -------------------------------------------------------------------------------- /src/g_typing_flags.ml4: -------------------------------------------------------------------------------- 1 | DECLARE PLUGIN "typing_flags_plugin" 2 | 3 | 4 | VERNAC COMMAND EXTEND TypingFlags1 CLASSIFIED AS SIDEFF 5 | | [ "Unset" "Guard" "Checking" ] -> [ Typing_flags_main.set_check_guarded false ] 6 | END 7 | 8 | VERNAC COMMAND EXTEND TypingFlags2 CLASSIFIED AS SIDEFF 9 | | [ "Set" "Guard" "Checking" ] -> [ Typing_flags_main.set_check_guarded true ] 10 | END 11 | 12 | VERNAC COMMAND EXTEND TypingFlags3 CLASSIFIED AS SIDEFF 13 | | [ "Set" "Type" "In" "Type" ] -> [ Typing_flags_main.set_check_universes false ] 14 | END 15 | 16 | VERNAC COMMAND EXTEND TypingFlags4 CLASSIFIED AS SIDEFF 17 | | [ "Unset" "Type" "In" "Type" ] -> [ Typing_flags_main.set_check_universes true ] 18 | END 19 | 20 | VERNAC COMMAND EXTEND TypingFlags5 CLASSIFIED AS QUERY 21 | | [ "Print" "Typing" "Flags" ] -> [ Typing_flags_main.print_typing_flags () ] 22 | END 23 | -------------------------------------------------------------------------------- /src/typing_flags_main.ml: -------------------------------------------------------------------------------- 1 | open Declarations 2 | open Environ 3 | open Pp 4 | 5 | let set_check_guarded b = 6 | let env = Global.env () in 7 | let flags = {(typing_flags env) with check_guarded = b} in 8 | Global.set_typing_flags flags 9 | 10 | let print_typing_flags () = 11 | let flags = typing_flags (Global.env ()) in 12 | Feedback.msg_notice (str "check_guarded: " ++ bool flags.check_guarded 13 | ++ str "\ncheck_universes: " ++ bool flags.check_universes) 14 | 15 | let set_check_universes b = 16 | let env = Global.env () in 17 | let flags = {(typing_flags env) with check_universes = b} in 18 | Global.set_typing_flags flags 19 | -------------------------------------------------------------------------------- /src/typing_flags_main.mli: -------------------------------------------------------------------------------- 1 | val set_check_guarded : bool -> unit 2 | 3 | val set_check_universes : bool -> unit 4 | 5 | val print_typing_flags : unit -> unit 6 | -------------------------------------------------------------------------------- /src/typing_flags_plugin.mlpack: -------------------------------------------------------------------------------- 1 | Typing_flags_main 2 | G_typing_flags -------------------------------------------------------------------------------- /theories/BreadthFirst.v: -------------------------------------------------------------------------------- 1 | (** Authors: 2 | Simon Boulier (INRIA) for the part before the def. of γ (without the simple lemmas) 3 | Ralph Matthes (IRIT - CNRS and Univ. of Toulouse) for the three different methods of verification, all suggested in 1995 4 | *) 5 | 6 | From TypingFlags Require Import Loader. 7 | 8 | 9 | Inductive tree := 10 | | leaf : nat -> tree 11 | | node : tree -> nat -> tree -> tree. 12 | 13 | Definition ex1 : tree. 14 | Proof. 15 | refine (node _ 1 (leaf 3)). 16 | refine (node _ 2 _). 17 | - exact (node (leaf 6) 4 (leaf 7)). 18 | - refine (node _ 5 (leaf 9)). 19 | exact (node (leaf 10) 8 (leaf 11)). 20 | Defined. 21 | 22 | Require Import List. 23 | Import ListNotations. 24 | 25 | Fixpoint zip (l1 l2 : list (list nat)) : list (list nat) := 26 | match l1, l2 with 27 | | [], l => l 28 | | l, [] => l 29 | | x :: l, x' :: l' => (x ++ x') :: (zip l l') 30 | end. 31 | 32 | Lemma zip_with_nil (l: list (list nat)): zip l [] = l. 33 | Proof. 34 | destruct l; reflexivity. 35 | Qed. 36 | 37 | Lemma zip_assoc (l1 l2 l3: list (list nat)): 38 | zip l1 (zip l2 l3) = zip (zip l1 l2) l3. 39 | Proof. 40 | revert l2 l3. 41 | induction l1. 42 | - reflexivity. 43 | - destruct l2. 44 | + reflexivity. 45 | + destruct l3. 46 | * reflexivity. 47 | * simpl. 48 | rewrite app_assoc. 49 | f_equal. 50 | apply IHl1. 51 | Qed. 52 | 53 | Fixpoint niveaux (t : tree) : list (list nat) := 54 | match t with 55 | | leaf x => [[x]] 56 | | node t1 x t2 => [x] :: (zip (niveaux t1) (niveaux t2)) 57 | end. 58 | 59 | Definition breadthfirst t := concat (niveaux t). 60 | 61 | Compute (niveaux ex1). 62 | Compute (breadthfirst ex1). 63 | 64 | Unset Guard Checking. 65 | Inductive Cor := 66 | | Over : Cor 67 | | Next : ((Cor -> list nat) -> list nat) -> Cor. 68 | Set Guard Checking. 69 | 70 | Definition apply (c : Cor) (k : Cor -> list nat) : list nat := 71 | match c with 72 | | Over => k Over 73 | | Next f => f k 74 | end. 75 | 76 | Fixpoint breadth (t : tree) (c : Cor) : Cor := 77 | match t with 78 | | leaf n => Next (fun k => n :: apply c k) 79 | | node l n r => Next (fun k => n :: apply c (fun c1 => k (breadth l (breadth r c1)))) 80 | end. 81 | 82 | Unset Guard Checking. 83 | Fixpoint ex (c : Cor) : list nat := 84 | match c with 85 | | Over => [] 86 | | Next f => f ex 87 | end. 88 | Set Guard Checking. 89 | 90 | Definition breadthfirst' t := ex (breadth t Over). 91 | 92 | Compute (breadth (leaf 3) Over). 93 | Compute (breadth (leaf 2) (breadth (leaf 3) Over)). 94 | Compute (breadthfirst' ex1). 95 | Print Assumptions breadthfirst'. 96 | 97 | (** ** Verification of the algorithm in Martin Hofmann's own style *) 98 | 99 | Fixpoint γ (ll:list (list nat))(c: Cor): Cor := 100 | match ll with 101 | | [] => c 102 | (* | [l] => Next ( fun k => l ++ apply c k ) *) 103 | | l :: ll1 => Next (fun k => l ++ apply c (fun c1 => k (γ ll1 c1)))end. 104 | 105 | Lemma MH_LemmaA (ll:list (list nat)): ex(γ ll Over) = concat ll. 106 | Proof. 107 | induction ll. 108 | - reflexivity. 109 | - simpl. 110 | f_equal. 111 | exact IHll. 112 | Qed. 113 | 114 | (** the following is required to get the proofs with Coq through although 115 | a pencil-and-paper proof is possible that all these equations hold 116 | w.r.t. convertibility for all input lists *) 117 | Require Import Logic.FunctionalExtensionality. 118 | 119 | Lemma MH_LemmaB (ll ll':list (list nat))(c: Cor): γ ll (γ ll' c) = γ (zip ll ll') c. 120 | Proof. 121 | revert ll' c. 122 | induction ll. 123 | - simpl. reflexivity. 124 | - induction ll'. 125 | + simpl. reflexivity. 126 | + simpl. 127 | intro c. 128 | f_equal. 129 | apply functional_extensionality; intro k. 130 | rewrite app_assoc. 131 | f_equal. 132 | f_equal. 133 | clear c IHll'. 134 | apply functional_extensionality; intro c. 135 | f_equal. 136 | apply IHll. 137 | Qed. 138 | 139 | Lemma MH_LemmaC (t: tree)(c: Cor): breadth t c = γ (niveaux t) c. 140 | Proof. 141 | revert c. 142 | induction t. 143 | - simpl. reflexivity. (* thanks to eta-equality *) 144 | - simpl. 145 | assert( aux : forall (c: Cor), breadth t1 (breadth t2 c) = γ (zip (niveaux t1) (niveaux t2)) c). 146 | { intro c. 147 | rewrite IHt2. 148 | rewrite IHt1. 149 | apply MH_LemmaB. 150 | } 151 | intro c. 152 | f_equal. 153 | apply functional_extensionality; intro k. 154 | f_equal. 155 | f_equal. 156 | clear c. 157 | apply functional_extensionality; intro c. 158 | f_equal. 159 | apply aux. 160 | Qed. 161 | 162 | Theorem MH_Verif (t: tree): breadthfirst' t = breadthfirst t. 163 | Proof. 164 | unfold breadthfirst', breadthfirst. 165 | rewrite MH_LemmaC. 166 | apply MH_LemmaA. 167 | Qed. 168 | 169 | Print Assumptions MH_Verif. 170 | 171 | 172 | (** ** Verification of the algorithm in the style of Ulrich Berger *) 173 | (** It does not need functional extensionality even in the Coq proof. *) 174 | 175 | Unset Guard Checking. 176 | Inductive rep: Cor -> list (list nat) -> Prop := 177 | | over: rep Over [] 178 | | next: forall (f: (Cor -> list nat) -> list nat)(l: list nat)(ll: list(list nat)), (forall (k: Cor -> list nat)(ll': list(list nat)), (forall (c: Cor)(ll1: list(list nat)), rep c ll1 -> k c = concat(zip ll' ll1)) -> f k = l ++ concat(zip ll' ll)) -> rep (Next f) (l::ll). 179 | (** is a non-strictly positive inductive proposition - could equivalently be defined impredicatively thanks to impredicativity of Prop *) 180 | 181 | Fixpoint rep_ind (R : Cor -> list (list nat) -> Prop)(HypO: R Over []) 182 | (HypN: forall (f: (Cor -> list nat) -> list nat)(l: list nat)(ll: list(list nat)), (forall (k: Cor -> list nat)(ll': list(list nat)), (forall (c: Cor)(ll1: list(list nat)), R c ll1 -> k c = concat(zip ll' ll1)) -> f k = l ++ concat(zip ll' ll)) -> R (Next f) (l::ll)) (c : Cor) (l : list (list nat))(Hyp: rep c l) {struct Hyp}: R c l := 183 | match Hyp in (rep c0 l0) return (R c0 l0) with 184 | | over => HypO 185 | | next f0 l0 ll0 Hyp0 => 186 | HypN _ _ _ (fun k ll' HypR => Hyp0 _ _ (fun c0 ll1 Hyp1 => HypR _ _ (rep_ind R HypO HypN c0 ll1 Hyp1))) 187 | end. 188 | (* interactively: 189 | Proof. 190 | refine (match Hyp in (rep c0 l0) return (R c0 l0) with 191 | | over => HypO 192 | | next f0 l0 ll0 Hyp0 => _ 193 | end). 194 | apply HypN. 195 | intros k ll' HypR. 196 | apply Hyp0. 197 | intros c0 ll1 Hyp1. 198 | apply HypR. 199 | apply (rep_ind R HypO HypN c0 ll1 Hyp1). 200 | Defined. 201 | *) 202 | Set Guard Checking. 203 | 204 | Definition isextractor (R: Cor -> list (list nat) -> Prop)(ll: list(list nat))(k:Cor -> list nat): Prop := 205 | forall (c: Cor)(ll1: list(list nat)), R c ll1 -> k c = concat(zip ll ll1). 206 | 207 | Check (next: forall (f: (Cor -> list nat) -> list nat)(l: list nat)(ll: list(list nat)), (forall (k: Cor -> list nat)(ll': list(list nat)), isextractor rep ll' k -> f k = l ++ concat(zip ll' ll)) -> rep (Next f) (l::ll)). 208 | 209 | Lemma UB_Lemma1: isextractor rep [] ex. 210 | Proof. 211 | red. 212 | intros ? ? Hyp. 213 | induction Hyp. 214 | - reflexivity. 215 | - simpl. 216 | change ll with (zip [] ll). 217 | apply H. 218 | intros ? ? Hyp. 219 | exact Hyp. 220 | Qed. 221 | 222 | Lemma UB_Lemma2 (t: tree)(c: Cor)(ll: list(list nat)): 223 | rep c ll -> rep (breadth t c) (zip (niveaux t) ll). 224 | Proof. 225 | revert c ll. 226 | induction t. 227 | - intros ? ?. intro Hyp; destruct Hyp. 228 | + simpl. 229 | apply next. 230 | intros ? ? Hyp1. 231 | simpl. 232 | f_equal. 233 | apply Hyp1. 234 | apply over. 235 | + simpl. 236 | apply next. 237 | intros ? ? Hyp1. 238 | simpl. 239 | f_equal. 240 | apply H. 241 | intros ? ? . 242 | apply Hyp1. 243 | - intros ? ?; intro Hyp; destruct Hyp. 244 | + simpl. 245 | apply next. 246 | intros ? ? Hyp1. 247 | simpl. 248 | f_equal. 249 | apply Hyp1. 250 | apply IHt1. 251 | replace (niveaux t2) with (zip (niveaux t2) []) by (apply zip_with_nil). 252 | apply IHt2. 253 | apply over. 254 | + simpl. 255 | apply next. 256 | intros ? ? Hyp1. 257 | simpl. 258 | f_equal. 259 | set (k0 := fun c1 : Cor => k (breadth t1 (breadth t2 c1))). 260 | assert (aux: isextractor rep (zip ll'(zip (niveaux t1) (niveaux t2))) k0). 261 | { red. 262 | intros ? ? Hyp2. 263 | unfold k0. 264 | set (aux2 := Hyp1 _ _ (IHt1 _ _ (IHt2 _ _ Hyp2))). 265 | rewrite aux2. 266 | f_equal. 267 | do 3 rewrite zip_assoc. 268 | reflexivity. 269 | } 270 | do 2 rewrite zip_assoc. 271 | apply H. 272 | intros ? ? Hyp2. 273 | red in aux. 274 | rewrite (aux _ _ Hyp2). 275 | f_equal. 276 | rewrite zip_assoc. 277 | reflexivity. 278 | Qed. 279 | 280 | Theorem UB_Verif (t: tree): breadthfirst' t = breadthfirst t. 281 | Proof. 282 | unfold breadthfirst', breadthfirst. 283 | set (aux := UB_Lemma2 t _ _ over). 284 | rewrite (UB_Lemma1 _ _ aux). 285 | rewrite zip_with_nil. 286 | reflexivity. 287 | Qed. 288 | 289 | Print Assumptions UB_Verif. 290 | 291 | (** ** Verification of the algorithm in the style of Anton Setzer *) 292 | 293 | Definition Cor' := list( list nat -> list nat). 294 | 295 | Fixpoint ϕ (c: Cor'): Cor := 296 | match c with 297 | | [] => Over 298 | | g::ll => Next(fun k: Cor -> list nat => g(k(ϕ ll))) 299 | end. 300 | 301 | (** "predicative" version of [breadth] *) 302 | Fixpoint breadthp (t: tree)(c: Cor'): Cor' := 303 | match t, c with 304 | | leaf n, [] => [cons n] 305 | | leaf n, g::l1 => (fun l:list nat => n::g l)::l1 306 | | node l n r, [] => (cons n) :: breadthp l (breadthp r []) 307 | | node l n r, g::l1 => (fun l:list nat => n::g l) :: breadthp l (breadthp r l1) 308 | end. 309 | 310 | Lemma AS_Lemma1 (t: tree)(c: Cor'): breadth t (ϕ c) = ϕ(breadthp t c). 311 | Proof. 312 | revert c. 313 | induction t; destruct c. 314 | - reflexivity. 315 | - reflexivity. 316 | - simpl. 317 | f_equal. 318 | apply functional_extensionality; intro k. 319 | do 2 f_equal. 320 | change Over with (ϕ []). 321 | rewrite (IHt2 []). 322 | apply IHt1. 323 | - simpl. 324 | f_equal. 325 | apply functional_extensionality; intro k. 326 | do 3 f_equal. 327 | rewrite IHt2. 328 | apply IHt1. 329 | Qed. 330 | 331 | Definition ψ (l: list nat): (list nat -> list nat) := fun l' => l ++ l'. 332 | Definition ψ': list(list nat) -> Cor' := map ψ. 333 | 334 | (** second refinement step for [breadthp] *) 335 | 336 | Fixpoint breadthp' (t: tree)(ll: list(list nat)): list(list nat) := 337 | match t, ll with 338 | | leaf n, [] => [[n]] 339 | | leaf n, l::ll => (n::l)::ll 340 | | node l n r, [] => [n] :: breadthp' l (breadthp' r []) 341 | | node l n r, l1::ll => (n::l1) :: breadthp' l (breadthp' r ll) 342 | end. 343 | 344 | Lemma AS_Lemma2 (t: tree)(ll: list(list nat)): breadthp t (ψ' ll) = ψ'(breadthp' t ll). 345 | Proof. 346 | revert ll. 347 | induction t; destruct ll. 348 | - reflexivity. 349 | - reflexivity. 350 | - simpl. 351 | f_equal. (* thanks to eta-equality *) 352 | change (@nil (forall _ : list nat, list nat)) with (ψ' []). 353 | rewrite (IHt2 []). 354 | apply IHt1. 355 | - simpl. 356 | f_equal. (* thanks to eta-equality *) 357 | rewrite IHt2. 358 | apply IHt1. 359 | Qed. 360 | 361 | Lemma AS_Lemma3 (t: tree)(ll: list(list nat)): breadthp' t ll = zip (niveaux t) ll. 362 | Proof. 363 | revert ll. 364 | induction t; destruct ll. 365 | - reflexivity. 366 | - reflexivity. 367 | - simpl. 368 | f_equal. 369 | rewrite IHt2. 370 | rewrite zip_with_nil. 371 | apply IHt1. 372 | - simpl. 373 | f_equal. 374 | rewrite IHt2. 375 | rewrite <- zip_assoc. 376 | apply IHt1. 377 | Qed. 378 | 379 | Lemma AS_Lemma4 (ll: list(list nat)): ex(ϕ(ψ' ll)) = concat ll. 380 | Proof. 381 | induction ll. 382 | - reflexivity. 383 | - simpl. 384 | rewrite IHll. 385 | reflexivity. 386 | Qed. 387 | 388 | Theorem AS_Verif (t: tree): breadthfirst' t = breadthfirst t. 389 | Proof. 390 | unfold breadthfirst', breadthfirst. 391 | change Over with (ϕ []). 392 | rewrite AS_Lemma1. 393 | change (@nil (forall _ : list nat, list nat)) with (ψ' []). 394 | rewrite AS_Lemma2. 395 | rewrite AS_Lemma3. 396 | rewrite zip_with_nil. 397 | apply AS_Lemma4. 398 | Qed. 399 | 400 | Print Assumptions AS_Verif. 401 | -------------------------------------------------------------------------------- /theories/Demo.v: -------------------------------------------------------------------------------- 1 | From TypingFlags Require Import Loader. 2 | 3 | Fail Definition T := let T := Type in (T : T). 4 | 5 | Set Type In Type. 6 | 7 | Print Typing Flags. 8 | 9 | Definition T := let T := Type in (T : T). 10 | 11 | 12 | 13 | 14 | Section t. 15 | 16 | Unset Guard Checking. 17 | 18 | Variable b : bool. 19 | 20 | Fixpoint f (n : nat) {struct n} : True := 21 | match n with 22 | | O => I 23 | | S _ => if b then f n else I 24 | end. 25 | 26 | Print Typing Flags. 27 | End t. 28 | 29 | 30 | Fail Definition g := Eval lazy in (f true). 31 | 32 | Unset Guard Checking. 33 | Definition g := Eval lazy in (f true). 34 | 35 | (* SetGuardChecking. *) 36 | Inductive T2 := 37 | l : (T2 -> False) -> T2. 38 | 39 | Goal False. 40 | assert T2. 41 | constructor. intro. 42 | remember H. 43 | destruct H. auto. 44 | remember H. 45 | destruct H. auto. 46 | (* Defined. (* Does not terminate? *) *) 47 | Qed. 48 | 49 | Check g. 50 | 51 | Compute (f false 12). 52 | 53 | Unset Type In Type. 54 | Set Guard Checking. 55 | Definition ff := f. 56 | Print Assumptions ff. 57 | -------------------------------------------------------------------------------- /theories/Loader.v: -------------------------------------------------------------------------------- 1 | Declare ML Module "typing_flags_plugin". -------------------------------------------------------------------------------- /theories/SameFringe.v: -------------------------------------------------------------------------------- 1 | (** Author: Ralph Matthes (IRIT - CNRS and Univ. of Toulouse) *) 2 | 3 | (** * Same Fringe problem solved with non-strictly positive inductive types *) 4 | 5 | (** The ideas for the following program have been communicated to me by Olivier Danvy on March 14, 2002. 6 | This was by way of an SML program. *) 7 | 8 | From TypingFlags Require Import Loader. 9 | Require Import Init.Nat. 10 | 11 | (** binary trees with only leaf labels *) 12 | Inductive tree := 13 | | leaf : nat -> tree 14 | | node : tree -> tree -> tree. 15 | 16 | Unset Guard Checking. 17 | Inductive Cor := 18 | | Over : Cor 19 | | Next : nat -> ((Cor -> bool) -> bool) -> Cor. 20 | Set Guard Checking. 21 | 22 | Unset Guard Checking. 23 | Fixpoint skim (c1 c2: Cor) {struct c1}: bool := 24 | match c1, c2 with 25 | | Over, Over => true 26 | | Over, Next n f => false 27 | | Next n f, Over => false 28 | | Next n1 f1, Next n2 f2 => if (negb (eqb n1 n2)) then false else f1 (fun c => f2 (skim c)) 29 | end. 30 | Set Guard Checking. 31 | 32 | Fixpoint walk (t: tree)(k: Cor -> bool)(f: (Cor -> bool) -> bool): bool:= 33 | match t with 34 | | leaf n => k (Next n f) 35 | | node l r => walk l k (fun k1 => walk r k1 f) 36 | end. 37 | 38 | Definition canf: (Cor -> bool) -> bool := fun k => k Over. 39 | 40 | Definition init (t: tree)(k: Cor -> bool): bool := walk t k canf. 41 | 42 | Definition smf (t1 t2: tree): bool := init t1 (fun c1 => init t2 (skim c1)). 43 | 44 | 45 | Definition ex1 := smf (node (leaf 1) (node (leaf 2) (leaf 3))) 46 | (node (node (leaf 1) (leaf 2)) (leaf 3)). 47 | Definition ex2 := smf (node (leaf 1) (node (leaf 2) (leaf 3))) 48 | (node (node (leaf 1) (leaf 2)) (leaf 4)). 49 | Definition ex3 := smf (node (leaf 1) (node (leaf 0) (leaf 3))) 50 | (node (node (leaf 1) (leaf 2)) (leaf 3)). 51 | Definition ex4 := smf (node (leaf 1) (node (leaf 2) (leaf 3))) 52 | (node (leaf 1) (node (leaf 2) (leaf 4))). 53 | 54 | 55 | Compute ex1. (* true *) 56 | Compute ex2. (* false *) 57 | Compute ex3. (* false *) 58 | Compute ex4. (* false *) 59 | 60 | Print Assumptions smf. 61 | --------------------------------------------------------------------------------